aboutsummaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorsacerdot2005-01-14 14:30:06 +0000
committersacerdot2005-01-14 14:30:06 +0000
commitb5657ff1939c6872ee3ccaeaf180a2f3da9e1876 (patch)
tree329838825dc03d431fc36489fdc818ca04d4bd2a /pretyping
parent6dbc9f181b90216958fd9d87f8426901b1e4c37e (diff)
Inductive.{type_of_inductive,type_of_constructor,arities_of_specif} changed
to accept a mind_specif (a couple mutual_inductive_body * one_inductive_body) instead of looking it up in the environment. A version of the same functions with the old type is put in Inductiveops (outside the kernel). git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@6589 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/inductiveops.ml18
-rw-r--r--pretyping/inductiveops.mli11
-rw-r--r--pretyping/retyping.ml1
-rw-r--r--pretyping/typing.ml1
4 files changed, 31 insertions, 0 deletions
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index cda1169dd7..74795f572c 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -18,6 +18,24 @@ open Declarations
open Environ
open Reductionops
+(* The following three functions are similar to the ones defined in
+ Inductive, but they expect an env *)
+
+let type_of_inductive env ind =
+ let specif = Inductive.lookup_mind_specif env ind in
+ Inductive.type_of_inductive specif
+
+(* Return type as quoted by the user *)
+let type_of_constructor env cstr =
+ let specif =
+ Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in
+ Inductive.type_of_constructor cstr specif
+
+(* Return constructor types in normal form *)
+let arities_of_constructors env ind =
+ let specif = Inductive.lookup_mind_specif env ind in
+ Inductive.arities_of_constructors ind specif
+
(* [inductive_family] = [inductive_instance] applied to global parameters *)
type inductive_family = inductive * constr list
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 9d67089f78..2bc8136787 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -14,6 +14,17 @@ open Declarations
open Environ
open Evd
+(* The following three functions are similar to the ones defined in
+ Inductive, but they expect an env *)
+
+val type_of_inductive : env -> inductive -> types
+
+(* Return type as quoted by the user *)
+val type_of_constructor : env -> constructor -> types
+
+(* Return constructor types in normal form *)
+val arities_of_constructors : env -> inductive -> types array
+
(* An inductive type with its parameters *)
type inductive_family
val make_ind_family : inductive * constr list -> inductive_family
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 5a71b28d8d..a75e48a8e9 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -11,6 +11,7 @@
open Util
open Term
open Inductive
+open Inductiveops
open Names
open Reductionops
open Environ
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 759a0c1a15..2f34957022 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -16,6 +16,7 @@ open Reductionops
open Type_errors
open Pretype_errors
open Inductive
+open Inductiveops
open Typeops
let meta_type env mv =