aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--engine/eConstr.ml6
-rw-r--r--engine/eConstr.mli6
-rw-r--r--engine/evarutil.mli3
-rw-r--r--engine/universes.mli4
-rw-r--r--kernel/constr.ml12
-rw-r--r--kernel/constr.mli22
-rw-r--r--kernel/term.ml11
-rw-r--r--kernel/term.mli13
-rw-r--r--pretyping/reductionops.mli2
9 files changed, 41 insertions, 38 deletions
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index 9026800f2..094841d69 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -17,11 +17,11 @@ open Evd
module API :
sig
type t
-val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term
-val kind_upto : Evd.evar_map -> constr -> (constr, types) Constr.kind_of_term
+val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term
+val kind_upto : Evd.evar_map -> constr -> (constr, types, Sorts.t, Univ.Instance.t) Constr.kind_of_term
val kind_of_type : Evd.evar_map -> t -> (t, t) kind_of_type
val whd_evar : Evd.evar_map -> t -> t
-val of_kind : (t, t) Constr.kind_of_term -> t
+val of_kind : (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term -> t
val of_constr : Constr.t -> t
val to_constr : evar_map -> t -> Constr.t
val unsafe_to_constr : t -> Constr.t
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
index 1ae71216f..e7287fc06 100644
--- a/engine/eConstr.mli
+++ b/engine/eConstr.mli
@@ -30,11 +30,11 @@ type rel_context = (constr, types) Context.Rel.pt
(** {5 Destructors} *)
-val kind : Evd.evar_map -> t -> (t, t) Constr.kind_of_term
+val kind : Evd.evar_map -> t -> (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term
(** Same as {!Constr.kind} except that it expands evars and normalizes
universes on the fly. *)
-val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t) Constr.kind_of_term
+val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t, Sorts.t, Univ.Instance.t) Constr.kind_of_term
val to_constr : Evd.evar_map -> t -> Constr.t
(** Returns the evar-normal form of the argument. See {!Evarutil.nf_evar}. *)
@@ -43,7 +43,7 @@ val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type
(** {5 Constructors} *)
-val of_kind : (t, t) Constr.kind_of_term -> t
+val of_kind : (t, t, Sorts.t, Univ.Instance.t) Constr.kind_of_term -> t
(** Construct a term from a view. *)
val of_constr : Constr.t -> t
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index da49d4e11..ca9591e71 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -178,7 +178,8 @@ val flush_and_check_evars : evar_map -> constr -> Constr.constr
(** Like {!Constr.kind} except that [kind_of_term sigma t] exposes [t]
as an evar [e] only if [e] is uninstantiated in [sigma]. Otherwise the
value of [e] in [sigma] is (recursively) used. *)
-val kind_of_term_upto : evar_map -> Constr.constr -> (Constr.constr,Constr.types) kind_of_term
+val kind_of_term_upto : evar_map -> Constr.constr ->
+ (Constr.constr, Constr.types, Sorts.t, Univ.Instance.t) kind_of_term
(** [eq_constr_univs_test sigma1 sigma2 t u] tests equality of [t] and
[u] up to existential variable instantiation and equalisable
diff --git a/engine/universes.mli b/engine/universes.mli
index c3e2055f3..932de941a 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -76,8 +76,8 @@ val eq_constr_univs_infer : UGraph.t -> 'a constraint_accumulator ->
{!eq_constr_univs_infer} taking kind-of-term functions, to expose
subterms of [m] and [n], arguments. *)
val eq_constr_univs_infer_with :
- (constr -> (constr,types) kind_of_term) ->
- (constr -> (constr,types) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
UGraph.t -> 'a constraint_accumulator -> constr -> constr -> 'a -> 'a option
(** [leq_constr_univs u a b] is [true, c] if [a] is convertible to [b]
diff --git a/kernel/constr.ml b/kernel/constr.ml
index 0ae3fb474..5a7561bf5 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -81,27 +81,27 @@ type pconstructor = constructor puniverses
(* [Var] is used for named variables and [Rel] for variables as
de Bruijn indices. *)
-type ('constr, 'types) kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of Sorts.t
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of pconstant
- | Ind of pinductive
- | Construct of pconstructor
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
| Proj of projection * 'constr
(* constr is the fixpoint of the previous type. Requires option
-rectypes of the Caml compiler to be set *)
-type t = (t,t) kind_of_term
+type t = (t, t, Sorts.t, Instance.t) kind_of_term
type constr = t
type existential = existential_key * constr array
diff --git a/kernel/constr.mli b/kernel/constr.mli
index c8be89fe2..700c235e6 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -188,7 +188,7 @@ type ('constr, 'types) pfixpoint =
type ('constr, 'types) pcofixpoint =
int * ('constr, 'types) prec_declaration
-type ('constr, 'types) kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Rel of int (** Gallina-variable introduced by [forall], [fun], [let-in], [fix], or [cofix]. *)
| Var of Id.t (** Gallina-variable that was introduced by Vernacular-command that extends
@@ -197,7 +197,7 @@ type ('constr, 'types) kind_of_term =
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of Sorts.t
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *)
| Lambda of Name.t * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *)
@@ -208,11 +208,11 @@ type ('constr, 'types) kind_of_term =
- [F] itself is not {!App}
- and [[|P1;..;Pn|]] is not empty. *)
- | Const of constant puniverses (** Gallina-variable that was introduced by Vernacular-command that extends the global environment
+ | Const of (constant * 'univs) (** Gallina-variable that was introduced by Vernacular-command that extends the global environment
(i.e. [Parameter], or [Axiom], or [Definition], or [Theorem] etc.) *)
- | Ind of inductive puniverses (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
- | Construct of constructor puniverses (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
+ | Ind of (inductive * 'univs) (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
+ | Construct of (constructor * 'univs) (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
@@ -222,8 +222,8 @@ type ('constr, 'types) kind_of_term =
least one argument and the function is not itself an applicative
term *)
-val kind : constr -> (constr, types) kind_of_term
-val of_kind : (constr, types) kind_of_term -> constr
+val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
+val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr
(** [equal a b] is true if [a] equals [b] modulo alpha, casts,
and application grouping *)
@@ -312,8 +312,8 @@ val compare_head_gen : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
constr -> constr -> bool
val compare_head_gen_leq_with :
- (constr -> (constr,types) kind_of_term) ->
- (constr -> (constr,types) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
(bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
(Sorts.t -> Sorts.t -> bool) ->
(constr -> constr -> bool) ->
@@ -325,8 +325,8 @@ val compare_head_gen_leq_with :
is used,rather than {!kind}, to expose the immediate subterms of
[c1] (resp. [c2]). *)
val compare_head_gen_with :
- (constr -> (constr,types) kind_of_term) ->
- (constr -> (constr,types) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
(bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
(Sorts.t -> Sorts.t -> bool) ->
(constr -> constr -> bool) ->
diff --git a/kernel/term.ml b/kernel/term.ml
index 3881cd12d..e5a681375 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -71,20 +71,21 @@ type pconstant = constant puniverses
type pinductive = inductive puniverses
type pconstructor = constructor puniverses
-type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
+ ('constr, 'types, 'sort, 'univs) Constr.kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of sorts
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of pconstant
- | Ind of pinductive
- | Construct of pconstructor
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
diff --git a/kernel/term.mli b/kernel/term.mli
index e393adb81..a9bb67705 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -68,20 +68,21 @@ type ('constr, 'types) prec_declaration =
type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint
type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint
-type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
+ ('constr, 'types, 'sort, 'univs) Constr.kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of sorts
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant puniverses
- | Ind of inductive puniverses
- | Construct of constructor puniverses
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
@@ -443,7 +444,7 @@ val leq_constr_univs : constr UGraph.check_function
application grouping and ignoring universe instances. *)
val eq_constr_nounivs : constr -> constr -> bool
-val kind_of_term : constr -> (constr, types) kind_of_term
+val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
(** Alias for [Constr.kind] *)
val constr_ord : constr -> constr -> int
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 15ddeb15c..18416b142 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -232,7 +232,7 @@ type 'a miota_args = {
val reducible_mind_case : evar_map -> constr -> bool
val reduce_mind_case : evar_map -> constr miota_args -> constr
-val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term
+val find_conclusion : env -> evar_map -> constr -> (constr, constr, Sorts.t, Univ.Instance.t) kind_of_term
val is_arity : env -> evar_map -> constr -> bool
val is_sort : env -> evar_map -> types -> bool