aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-11-06 23:27:09 +0100
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-11-06 23:46:52 +0100
commit03e21974a3e971a294533bffb81877dc1bd270b6 (patch)
tree1b37339378f6bc93288b61f707efb6b08f992dc5
parentf3abbc55ef160d1a65d4467bfe9b25b30b965a46 (diff)
[api] Move structures deprecated in the API to the core.
We do up to `Term` which is the main bulk of the changes.
-rw-r--r--API/API.mli11
-rw-r--r--dev/top_printers.ml7
-rw-r--r--engine/eConstr.ml11
-rw-r--r--engine/eConstr.mli2
-rw-r--r--engine/evarutil.ml15
-rw-r--r--engine/evarutil.mli6
-rw-r--r--engine/evd.ml19
-rw-r--r--engine/evd.mli56
-rw-r--r--engine/namegen.mli3
-rw-r--r--engine/termops.ml9
-rw-r--r--engine/termops.mli8
-rw-r--r--engine/uState.ml4
-rw-r--r--engine/uState.mli14
-rw-r--r--engine/universes.ml25
-rw-r--r--engine/universes.mli76
-rw-r--r--interp/constrextern.ml2
-rw-r--r--interp/constrextern.mli3
-rw-r--r--interp/constrintern.mli2
-rw-r--r--interp/declare.ml6
-rw-r--r--interp/declare.mli6
-rw-r--r--interp/discharge.ml1
-rw-r--r--interp/impargs.ml13
-rw-r--r--interp/impargs.mli2
-rw-r--r--interp/notation.ml4
-rw-r--r--interp/notation.mli4
-rw-r--r--intf/constrexpr.ml2
-rw-r--r--intf/glob_term.ml2
-rw-r--r--intf/misctypes.ml3
-rw-r--r--intf/notation_term.ml2
-rw-r--r--intf/pattern.ml4
-rw-r--r--kernel/cClosure.ml12
-rw-r--r--kernel/cClosure.mli2
-rw-r--r--kernel/cbytecodes.ml11
-rw-r--r--kernel/cbytecodes.mli8
-rw-r--r--kernel/cbytegen.ml32
-rw-r--r--kernel/cbytegen.mli10
-rw-r--r--kernel/constr.mli2
-rw-r--r--kernel/cooking.ml23
-rw-r--r--kernel/cooking.mli4
-rw-r--r--kernel/csymtable.ml4
-rw-r--r--kernel/csymtable.mli2
-rw-r--r--kernel/declarations.ml20
-rw-r--r--kernel/declareops.ml14
-rw-r--r--kernel/declareops.mli4
-rw-r--r--kernel/entries.ml14
-rw-r--r--kernel/environ.ml18
-rw-r--r--kernel/environ.mli12
-rw-r--r--kernel/indtypes.ml35
-rw-r--r--kernel/indtypes.mli2
-rw-r--r--kernel/inductive.ml115
-rw-r--r--kernel/inductive.mli26
-rw-r--r--kernel/mod_subst.ml4
-rw-r--r--kernel/mod_subst.mli7
-rw-r--r--kernel/modops.ml4
-rw-r--r--kernel/modops.mli2
-rw-r--r--kernel/names.mli1
-rw-r--r--kernel/nativecode.ml12
-rw-r--r--kernel/nativecode.mli4
-rw-r--r--kernel/nativeconv.mli2
-rw-r--r--kernel/nativeinstr.mli4
-rw-r--r--kernel/nativelambda.ml14
-rw-r--r--kernel/nativelambda.mli2
-rw-r--r--kernel/nativevalues.ml10
-rw-r--r--kernel/nativevalues.mli6
-rw-r--r--kernel/opaqueproof.ml6
-rw-r--r--kernel/opaqueproof.mli16
-rw-r--r--kernel/pre_env.ml2
-rw-r--r--kernel/pre_env.mli2
-rw-r--r--kernel/reduction.ml44
-rw-r--r--kernel/reduction.mli8
-rw-r--r--kernel/retroknowledge.ml6
-rw-r--r--kernel/retroknowledge.mli6
-rw-r--r--kernel/safe_typing.mli20
-rw-r--r--kernel/sorts.ml4
-rw-r--r--kernel/sorts.mli6
-rw-r--r--kernel/subtyping.ml13
-rw-r--r--kernel/term.ml2
-rw-r--r--kernel/term.mli278
-rw-r--r--kernel/term_typing.ml16
-rw-r--r--kernel/term_typing.mli4
-rw-r--r--kernel/type_errors.ml7
-rw-r--r--kernel/type_errors.mli12
-rw-r--r--kernel/typeops.ml15
-rw-r--r--kernel/typeops.mli10
-rw-r--r--kernel/uGraph.ml1
-rw-r--r--kernel/uGraph.mli63
-rw-r--r--kernel/univ.mli200
-rw-r--r--kernel/vars.mli4
-rw-r--r--kernel/vconv.mli2
-rw-r--r--kernel/vm.ml11
-rw-r--r--kernel/vm.mli18
-rw-r--r--library/declaremods.mli2
-rw-r--r--library/decls.ml2
-rw-r--r--library/decls.mli4
-rw-r--r--library/global.ml2
-rw-r--r--library/global.mli18
-rw-r--r--library/globnames.ml6
-rw-r--r--library/globnames.mli2
-rw-r--r--library/heads.ml3
-rw-r--r--library/heads.mli2
-rw-r--r--library/lib.ml4
-rw-r--r--library/lib.mli6
-rw-r--r--library/library.ml10
-rw-r--r--library/library.mli4
-rw-r--r--library/nameops.mli4
-rw-r--r--library/univops.ml8
-rw-r--r--library/univops.mli6
-rw-r--r--plugins/btauto/refl_btauto.ml24
-rw-r--r--plugins/cc/ccalgo.ml13
-rw-r--r--plugins/cc/ccalgo.mli4
-rw-r--r--plugins/cc/ccproof.ml2
-rw-r--r--plugins/cc/ccproof.mli2
-rw-r--r--plugins/cc/cctac.ml7
-rw-r--r--plugins/derive/derive.ml5
-rw-r--r--plugins/extraction/extract_env.ml8
-rw-r--r--plugins/extraction/extract_env.mli2
-rw-r--r--plugins/extraction/extraction.ml31
-rw-r--r--plugins/extraction/extraction.mli2
-rw-r--r--plugins/extraction/table.mli2
-rw-r--r--plugins/firstorder/formula.ml4
-rw-r--r--plugins/firstorder/formula.mli2
-rw-r--r--plugins/firstorder/instances.ml2
-rw-r--r--plugins/firstorder/rules.mli4
-rw-r--r--plugins/firstorder/sequent.ml12
-rw-r--r--plugins/firstorder/sequent.mli6
-rw-r--r--plugins/firstorder/unify.mli2
-rw-r--r--plugins/fourier/fourierR.ml24
-rw-r--r--plugins/funind/functional_principles_types.ml45
-rw-r--r--plugins/funind/functional_principles_types.mli2
-rw-r--r--plugins/funind/glob_term_to_relation.ml5
-rw-r--r--plugins/funind/glob_term_to_relation.mli2
-rw-r--r--plugins/funind/indfun_common.ml10
-rw-r--r--plugins/funind/indfun_common.mli8
-rw-r--r--plugins/funind/invfun.ml1
-rw-r--r--plugins/funind/merge.ml13
-rw-r--r--plugins/funind/recdef.ml18
-rw-r--r--plugins/funind/recdef.mli7
-rw-r--r--plugins/ltac/extratactics.ml46
-rw-r--r--plugins/ltac/g_class.ml42
-rw-r--r--plugins/ltac/pptactic.ml8
-rw-r--r--plugins/ltac/rewrite.ml8
-rw-r--r--plugins/ltac/rewrite.mli2
-rw-r--r--plugins/ltac/taccoerce.ml6
-rw-r--r--plugins/micromega/coq_micromega.ml7
-rw-r--r--plugins/nsatz/nsatz.ml30
-rw-r--r--plugins/nsatz/nsatz.mli2
-rw-r--r--plugins/quote/quote.ml12
-rw-r--r--plugins/romega/const_omega.ml104
-rw-r--r--plugins/romega/const_omega.mli155
-rw-r--r--plugins/romega/refl_omega.ml19
-rw-r--r--plugins/rtauto/refl_tauto.ml3
-rw-r--r--plugins/rtauto/refl_tauto.mli2
-rw-r--r--plugins/setoid_ring/newring.ml18
-rw-r--r--plugins/setoid_ring/newring_ast.mli2
-rw-r--r--plugins/ssr/ssrcommon.ml34
-rw-r--r--plugins/ssr/ssrcommon.mli6
-rw-r--r--plugins/ssr/ssrequality.ml5
-rw-r--r--plugins/ssr/ssrfwd.ml11
-rw-r--r--plugins/ssr/ssrvernac.ml45
-rw-r--r--plugins/ssrmatching/ssrmatching.ml481
-rw-r--r--plugins/ssrmatching/ssrmatching.mli2
-rw-r--r--pretyping/arguments_renaming.ml3
-rw-r--r--pretyping/arguments_renaming.mli2
-rw-r--r--pretyping/cases.ml4
-rw-r--r--pretyping/cases.mli2
-rw-r--r--pretyping/cbv.ml12
-rw-r--r--pretyping/cbv.mli2
-rw-r--r--pretyping/classops.ml4
-rw-r--r--pretyping/constr_matching.mli2
-rw-r--r--pretyping/detyping.ml2
-rw-r--r--pretyping/detyping.mli3
-rw-r--r--pretyping/evarconv.ml5
-rw-r--r--pretyping/evarconv.mli2
-rw-r--r--pretyping/evardefine.ml3
-rw-r--r--pretyping/evardefine.mli3
-rw-r--r--pretyping/evarsolve.ml3
-rw-r--r--pretyping/evarsolve.mli2
-rw-r--r--pretyping/indrec.ml17
-rw-r--r--pretyping/indrec.mli20
-rw-r--r--pretyping/inductiveops.ml5
-rw-r--r--pretyping/inductiveops.mli14
-rw-r--r--pretyping/nativenorm.ml7
-rw-r--r--pretyping/patternops.ml9
-rw-r--r--pretyping/pretype_errors.ml2
-rw-r--r--pretyping/pretype_errors.mli6
-rw-r--r--pretyping/pretyping.mli2
-rw-r--r--pretyping/recordops.ml14
-rw-r--r--pretyping/recordops.mli4
-rw-r--r--pretyping/reductionops.ml4
-rw-r--r--pretyping/reductionops.mli2
-rw-r--r--pretyping/retyping.ml11
-rw-r--r--pretyping/retyping.mli7
-rw-r--r--pretyping/typeclasses.ml1
-rw-r--r--pretyping/typeclasses.mli2
-rw-r--r--pretyping/typing.ml6
-rw-r--r--pretyping/typing.mli4
-rw-r--r--pretyping/unification.ml10
-rw-r--r--pretyping/unification.mli2
-rw-r--r--pretyping/vnorm.ml5
-rw-r--r--printing/printer.ml1
-rw-r--r--printing/printer.mli12
-rw-r--r--printing/printmod.ml4
-rw-r--r--proofs/clenv.mli2
-rw-r--r--proofs/goal.ml2
-rw-r--r--proofs/logic.ml23
-rw-r--r--proofs/logic.mli2
-rw-r--r--proofs/pfedit.mli2
-rw-r--r--proofs/proof_global.ml2
-rw-r--r--proofs/proof_global.mli2
-rw-r--r--proofs/proof_type.ml2
-rw-r--r--proofs/redexpr.mli2
-rw-r--r--proofs/refine.mli2
-rw-r--r--proofs/refiner.mli2
-rw-r--r--proofs/tacmach.mli2
-rw-r--r--stm/asyncTaskQueue.ml2
-rw-r--r--tactics/autorewrite.ml3
-rw-r--r--tactics/autorewrite.mli6
-rw-r--r--tactics/elimschemes.ml3
-rw-r--r--tactics/elimschemes.mli2
-rw-r--r--tactics/eqschemes.ml7
-rw-r--r--tactics/eqschemes.mli12
-rw-r--r--tactics/hints.ml4
-rw-r--r--tactics/hints.mli8
-rw-r--r--tactics/ind_tables.ml2
-rw-r--r--tactics/ind_tables.mli2
-rw-r--r--tactics/tacticals.ml9
-rw-r--r--tactics/tacticals.mli14
-rw-r--r--tactics/tactics.ml11
-rw-r--r--tactics/tactics.mli2
-rw-r--r--tactics/term_dnet.ml6
-rw-r--r--tactics/term_dnet.mli2
-rw-r--r--vernac/assumptions.ml6
-rw-r--r--vernac/assumptions.mli4
-rw-r--r--vernac/auto_ind_decl.ml3
-rw-r--r--vernac/class.ml3
-rw-r--r--vernac/classes.ml1
-rw-r--r--vernac/classes.mli2
-rw-r--r--vernac/command.ml5
-rw-r--r--vernac/command.mli2
-rw-r--r--vernac/declareDef.mli2
-rw-r--r--vernac/indschemes.ml3
-rw-r--r--vernac/indschemes.mli2
-rw-r--r--vernac/lemmas.ml9
-rw-r--r--vernac/lemmas.mli2
-rw-r--r--vernac/obligations.ml33
-rw-r--r--vernac/obligations.mli12
-rw-r--r--vernac/record.ml5
-rw-r--r--vernac/record.mli2
-rw-r--r--vernac/search.ml2
-rw-r--r--vernac/search.mli2
250 files changed, 1502 insertions, 1385 deletions
diff --git a/API/API.mli b/API/API.mli
index ccb71179d..065caeac3 100644
--- a/API/API.mli
+++ b/API/API.mli
@@ -553,6 +553,17 @@ sig
| CoFix of ('constr, 'types) pcofixpoint
| Proj of Projection.t * '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
+
+val map_with_binders :
+ ('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
+val map : (constr -> constr) -> constr -> constr
+
+val fold : ('a -> constr -> 'a) -> 'a -> constr -> 'a
+val iter : (constr -> unit) -> constr -> unit
+val compare_head : (constr -> constr -> bool) -> constr -> constr -> bool
+
val equal : t -> t -> bool
val eq_constr_nounivs : t -> t -> bool
val compare : t -> t -> int
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 0d833d33b..2dae96575 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -8,6 +8,7 @@
(* Printers for the ocaml toplevel. *)
+open Sorts
open Util
open Pp
open Names
@@ -17,7 +18,7 @@ open Nameops
open Univ
open Environ
open Printer
-open Term
+open Constr
open Evd
open Goptions
open Genarg
@@ -242,7 +243,7 @@ let cast_kind_display k =
| NATIVEcast -> "NATIVEcast"
let constr_display csr =
- let rec term_display c = match kind_of_term c with
+ let rec term_display c = match kind c with
| Rel n -> "Rel("^(string_of_int n)^")"
| Meta n -> "Meta("^(string_of_int n)^")"
| Var id -> "Var("^(Id.to_string id)^")"
@@ -314,7 +315,7 @@ let constr_display csr =
open Format;;
let print_pure_constr csr =
- let rec term_display c = match kind_of_term c with
+ let rec term_display c = match Constr.kind c with
| Rel n -> print_string "#"; print_int n
| Meta n -> print_string "Meta("; print_int n; print_string ")"
| Var id -> print_string (Id.to_string id)
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
index a54c08297..bcfbc8081 100644
--- a/engine/eConstr.ml
+++ b/engine/eConstr.ml
@@ -10,6 +10,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Context
open Evd
@@ -34,7 +35,7 @@ end
type t
val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.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 kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type
val whd_evar : Evd.evar_map -> t -> t
val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t
val of_constr : Constr.t -> t
@@ -84,16 +85,16 @@ let rec whd_evar sigma c =
| Some c -> whd_evar sigma c
| None -> c
end
- | App (f, args) when isEvar f ->
+ | App (f, args) when Term.isEvar f ->
(** Enforce smart constructor invariant on applications *)
- let ev = destEvar f in
+ let ev = Term.destEvar f in
begin match safe_evar_value sigma ev with
| None -> c
| Some f -> whd_evar sigma (mkApp (f, args))
end
- | Cast (c0, k, t) when isEvar c0 ->
+ | Cast (c0, k, t) when Term.isEvar c0 ->
(** Enforce smart constructor invariant on casts. *)
- let ev = destEvar c0 in
+ let ev = Term.destEvar c0 in
begin match safe_evar_value sigma ev with
| None -> c
| Some c -> whd_evar sigma (mkCast (c, k, t))
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
index a4ad233b0..e9ef302cf 100644
--- a/engine/eConstr.mli
+++ b/engine/eConstr.mli
@@ -93,7 +93,7 @@ val mkEvar : t pexistential -> t
val mkSort : Sorts.t -> t
val mkProp : t
val mkSet : t
-val mkType : Univ.universe -> t
+val mkType : Univ.Universe.t -> t
val mkCast : t * cast_kind * t -> t
val mkProd : Name.t * t * t -> t
val mkLambda : Name.t * t * t -> t
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 38efcca05..df4ef2ce7 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -10,6 +10,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Termops
open Namegen
open Pre_env
@@ -56,12 +57,12 @@ let new_global evd x =
exception Uninstantiated_evar of existential_key
let rec flush_and_check_evars sigma c =
- match kind_of_term c with
+ match kind c with
| Evar (evk,_ as ev) ->
(match existential_opt_value sigma ev with
| None -> raise (Uninstantiated_evar evk)
| Some c -> flush_and_check_evars sigma c)
- | _ -> map_constr (flush_and_check_evars sigma) c
+ | _ -> Constr.map (flush_and_check_evars sigma) c
let flush_and_check_evars sigma c =
flush_and_check_evars sigma (EConstr.Unsafe.to_constr c)
@@ -162,7 +163,7 @@ exception NoHeadEvar
let head_evar sigma c =
(** FIXME: this breaks if using evar-insensitive code *)
let c = EConstr.Unsafe.to_constr c in
- let rec hrec c = match kind_of_term c with
+ let rec hrec c = match kind c with
| Evar (evk,_) -> evk
| Case (_,_,c,_) -> hrec c
| App (c,_) -> hrec c
@@ -485,7 +486,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
(ie the hypotheses ids have been removed from the contexts of
evars). [global] should be true iff there is some variable of [ids] which
is a section variable *)
- match kind_of_term c with
+ match kind c with
| Var id' ->
if Id.Set.mem id' ids then raise (ClearDependencyError (id', err)) else c
@@ -552,7 +553,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
evdref := evd;
Evd.existential_value !evdref ev
- | _ -> map_constr (check_and_clear_in_constr env evdref err ids global) c
+ | _ -> Constr.map (check_and_clear_in_constr env evdref err ids global) c
let clear_hyps_in_evi_main env evdref hyps terms ids =
(* clear_hyps_in_evi erases hypotheses ids in hyps, checking if some
@@ -696,10 +697,10 @@ let undefined_evars_of_evar_info evd evi =
do not have this luxury, and need the more complete version. *)
let occur_evar_upto sigma n c =
let c = EConstr.Unsafe.to_constr c in
- let rec occur_rec c = match kind_of_term c with
+ let rec occur_rec c = match kind c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar e -> Option.iter occur_rec (existential_opt_value sigma e)
- | _ -> iter_constr occur_rec c
+ | _ -> Constr.iter occur_rec c
in
try occur_rec c; false with Occur -> true
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index 2f85bc733..62288ced4 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Evd
open Environ
open EConstr
@@ -54,11 +54,11 @@ val e_new_evar :
val new_type_evar :
env -> evar_map -> ?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid ->
- evar_map * (constr * sorts)
+ evar_map * (constr * Sorts.t)
val e_new_type_evar : env -> evar_map ref ->
?src:Evar_kinds.t Loc.located -> ?filter:Filter.t ->
- ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * sorts
+ ?naming:Misctypes.intro_pattern_naming_expr -> ?principal:bool -> rigid -> constr * Sorts.t
val new_Type : ?rigid:rigid -> env -> evar_map -> evar_map * constr
val e_new_Type : ?rigid:rigid -> env -> evar_map ref -> constr
diff --git a/engine/evd.ml b/engine/evd.ml
index 86ab2263f..a1cb0ec68 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -8,10 +8,11 @@
open Pp
open CErrors
+open Sorts
open Util
open Names
open Nameops
-open Term
+open Constr
open Vars
open Environ
@@ -126,7 +127,7 @@ end
module Store = Store.Make ()
-type evar = Term.existential_key
+type evar = existential_key
let string_of_existential evk = "?X" ^ string_of_int (Evar.repr evk)
@@ -242,7 +243,7 @@ let evar_instance_array test_id info args =
instrec filter (evar_context info) 0
let make_evar_instance_array info args =
- evar_instance_array (NamedDecl.get_id %> isVarId) info args
+ evar_instance_array (NamedDecl.get_id %> Term.isVarId) info args
let instantiate_evar_array info c args =
let inst = make_evar_instance_array info args in
@@ -280,9 +281,9 @@ type 'a freelisted = {
(* Collects all metavars appearing in a constr *)
let metavars_of c =
let rec collrec acc c =
- match kind_of_term c with
+ match kind c with
| Meta mv -> Int.Set.add mv acc
- | _ -> Term.fold_constr collrec acc c
+ | _ -> Constr.fold collrec acc c
in
collrec Int.Set.empty c
@@ -706,10 +707,10 @@ let extract_all_conv_pbs evd =
extract_conv_pbs evd (fun _ -> true)
let loc_of_conv_pb evd (pbty,env,t1,t2) =
- match kind_of_term (fst (decompose_app t1)) with
+ match kind (fst (Term.decompose_app t1)) with
| Evar (evk1,_) -> fst (evar_source evk1 evd)
| _ ->
- match kind_of_term (fst (decompose_app t2)) with
+ match kind (fst (Term.decompose_app t2)) with
| Evar (evk2,_) -> fst (evar_source evk2 evd)
| _ -> None
@@ -720,9 +721,9 @@ let loc_of_conv_pb evd (pbty,env,t1,t2) =
let evars_of_term c =
let rec evrec acc c =
- match kind_of_term c with
+ match kind c with
| Evar (n, l) -> Evar.Set.add n (Array.fold_left evrec acc l)
- | _ -> Term.fold_constr evrec acc c
+ | _ -> Constr.fold evrec acc c
in
evrec Evar.Set.empty c
diff --git a/engine/evd.mli b/engine/evd.mli
index 7576b3d5f..45ca1a365 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -9,7 +9,7 @@
open Util
open Loc
open Names
-open Term
+open Constr
open Environ
(** This file defines the pervasive unification state used everywhere in Coq
@@ -488,10 +488,10 @@ val univ_flexible_alg : rigid
type 'a in_evar_universe_context = 'a * evar_universe_context
-val evar_universe_context_set : evar_universe_context -> Univ.universe_context_set
+val evar_universe_context_set : evar_universe_context -> Univ.ContextSet.t
val evar_universe_context_constraints : evar_universe_context -> Univ.constraints
-val evar_context_universe_context : evar_universe_context -> Univ.universe_context
-val evar_universe_context_of : Univ.universe_context_set -> evar_universe_context
+val evar_context_universe_context : evar_universe_context -> Univ.UContext.t
+val evar_universe_context_of : Univ.ContextSet.t -> evar_universe_context
val empty_evar_universe_context : evar_universe_context
val union_evar_universe_context : evar_universe_context -> evar_universe_context ->
evar_universe_context
@@ -503,10 +503,10 @@ val evar_universe_context_of_binders :
Universes.universe_binders -> evar_universe_context
val make_evar_universe_context : env -> (Id.t located) list option -> evar_universe_context
-val restrict_universe_context : evar_map -> Univ.universe_set -> evar_map
+val restrict_universe_context : evar_map -> Univ.LSet.t -> evar_map
(** Raises Not_found if not a name for a universe in this map. *)
-val universe_of_name : evar_map -> string -> Univ.universe_level
-val add_universe_name : evar_map -> string -> Univ.universe_level -> evar_map
+val universe_of_name : evar_map -> string -> Univ.Level.t
+val add_universe_name : evar_map -> string -> Univ.Level.t -> evar_map
val add_constraints_context : evar_universe_context ->
Univ.constraints -> evar_universe_context
@@ -518,50 +518,50 @@ val normalize_evar_universe_context_variables : evar_universe_context ->
val normalize_evar_universe_context : evar_universe_context ->
evar_universe_context
-val new_univ_level_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Univ.universe_level
-val new_univ_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Univ.universe
-val new_sort_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * sorts
+val new_univ_level_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Univ.Level.t
+val new_univ_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Univ.Universe.t
+val new_sort_variable : ?loc:Loc.t -> ?name:string -> rigid -> evar_map -> evar_map * Sorts.t
val add_global_univ : evar_map -> Univ.Level.t -> evar_map
val universe_rigidity : evar_map -> Univ.Level.t -> rigid
-val make_flexible_variable : evar_map -> algebraic:bool -> Univ.universe_level -> evar_map
+val make_flexible_variable : evar_map -> algebraic:bool -> Univ.Level.t -> evar_map
(** See [UState.make_flexible_variable] *)
-val is_sort_variable : evar_map -> sorts -> Univ.universe_level option
+val is_sort_variable : evar_map -> Sorts.t -> Univ.Level.t option
(** [is_sort_variable evm s] returns [Some u] or [None] if [s] is
not a local sort variable declared in [evm] *)
val is_flexible_level : evar_map -> Univ.Level.t -> bool
-(* val normalize_universe_level : evar_map -> Univ.universe_level -> Univ.universe_level *)
-val normalize_universe : evar_map -> Univ.universe -> Univ.universe
-val normalize_universe_instance : evar_map -> Univ.universe_instance -> Univ.universe_instance
+(* val normalize_universe_level : evar_map -> Univ.Level.t -> Univ.Level.t *)
+val normalize_universe : evar_map -> Univ.Universe.t -> Univ.Universe.t
+val normalize_universe_instance : evar_map -> Univ.Instance.t -> Univ.Instance.t
-val set_leq_sort : env -> evar_map -> sorts -> sorts -> evar_map
-val set_eq_sort : env -> evar_map -> sorts -> sorts -> evar_map
-val has_lub : evar_map -> Univ.universe -> Univ.universe -> evar_map
-val set_eq_level : evar_map -> Univ.universe_level -> Univ.universe_level -> evar_map
-val set_leq_level : evar_map -> Univ.universe_level -> Univ.universe_level -> evar_map
+val set_leq_sort : env -> evar_map -> Sorts.t -> Sorts.t -> evar_map
+val set_eq_sort : env -> evar_map -> Sorts.t -> Sorts.t -> evar_map
+val has_lub : evar_map -> Univ.Universe.t -> Univ.Universe.t -> evar_map
+val set_eq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_map
+val set_leq_level : evar_map -> Univ.Level.t -> Univ.Level.t -> evar_map
val set_eq_instances : ?flex:bool ->
- evar_map -> Univ.universe_instance -> Univ.universe_instance -> evar_map
+ evar_map -> Univ.Instance.t -> Univ.Instance.t -> evar_map
-val check_eq : evar_map -> Univ.universe -> Univ.universe -> bool
-val check_leq : evar_map -> Univ.universe -> Univ.universe -> bool
+val check_eq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool
+val check_leq : evar_map -> Univ.Universe.t -> Univ.Universe.t -> bool
val evar_universe_context : evar_map -> evar_universe_context
-val universe_context_set : evar_map -> Univ.universe_context_set
+val universe_context_set : evar_map -> Univ.ContextSet.t
val universe_context : names:(Id.t located) list -> extensible:bool -> evar_map ->
- (Id.t * Univ.Level.t) list * Univ.universe_context
+ (Id.t * Univ.Level.t) list * Univ.UContext.t
val universe_subst : evar_map -> Universes.universe_opt_subst
val universes : evar_map -> UGraph.t
val check_univ_decl : evar_map -> UState.universe_decl ->
- Universes.universe_binders * Univ.universe_context
+ Universes.universe_binders * Univ.UContext.t
val merge_universe_context : evar_map -> evar_universe_context -> evar_map
val set_universe_context : evar_map -> evar_universe_context -> evar_map
-val merge_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> Univ.universe_context_set -> evar_map
+val merge_context_set : ?loc:Loc.t -> ?sideff:bool -> rigid -> evar_map -> Univ.ContextSet.t -> evar_map
val merge_universe_subst : evar_map -> Universes.universe_opt_subst -> evar_map
val with_context_set : ?loc:Loc.t -> rigid -> evar_map -> 'a Univ.in_universe_context_set -> evar_map * 'a
@@ -579,7 +579,7 @@ val update_sigma_env : evar_map -> env -> evar_map
(** Polymorphic universes *)
-val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> env -> evar_map -> sorts_family -> evar_map * sorts
+val fresh_sort_in_family : ?loc:Loc.t -> ?rigid:rigid -> env -> evar_map -> Sorts.family -> evar_map * Sorts.t
val fresh_constant_instance : ?loc:Loc.t -> env -> evar_map -> Constant.t -> evar_map * pconstant
val fresh_inductive_instance : ?loc:Loc.t -> env -> evar_map -> inductive -> evar_map * pinductive
val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor -> evar_map * pconstructor
diff --git a/engine/namegen.mli b/engine/namegen.mli
index d29b69259..abeed9f62 100644
--- a/engine/namegen.mli
+++ b/engine/namegen.mli
@@ -9,7 +9,6 @@
(** This file features facilities to generate fresh names. *)
open Names
-open Term
open Environ
open Evd
open EConstr
@@ -27,7 +26,7 @@ val default_dependent_ident : Id.t (* "x" *)
Generating "intuitive" names from their type *)
val lowercase_first_char : Id.t -> string
-val sort_hdchar : sorts -> string
+val sort_hdchar : Sorts.t -> string
val hdchar : env -> evar_map -> types -> string
val id_of_name_using_hdchar : env -> evar_map -> types -> Name.t -> Id.t
val named_hd : env -> evar_map -> types -> Name.t -> Name.t
diff --git a/engine/termops.ml b/engine/termops.ml
index 67533cce4..78dbdb11a 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -12,6 +12,7 @@ open Util
open Names
open Nameops
open Term
+open Constr
open Vars
open Environ
@@ -46,7 +47,7 @@ let pr_puniverses p u =
if Univ.Instance.is_empty u then p
else p ++ str"(*" ++ Univ.Instance.pr Universes.pr_with_global_universes u ++ str"*)"
-let rec pr_constr c = match kind_of_term c with
+let rec pr_constr c = match kind c with
| Rel n -> str "#"++int n
| Meta n -> str "Meta(" ++ int n ++ str ")"
| Var id -> pr_id id
@@ -798,7 +799,7 @@ let fold_constr_with_binders sigma g f n acc c =
let iter_constr_with_full_binders g f l c =
let open RelDecl in
- match kind_of_term c with
+ match kind c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> ()
| Cast (c,_, t) -> f l c; f l t
@@ -983,9 +984,9 @@ let isMetaOf sigma mv c =
match EConstr.kind sigma c with Meta mv' -> Int.equal mv mv' | _ -> false
let rec subst_meta bl c =
- match kind_of_term c with
+ match kind c with
| Meta i -> (try Int.List.assoc i bl with Not_found -> c)
- | _ -> map_constr (subst_meta bl) c
+ | _ -> Constr.map (subst_meta bl) c
let rec strip_outer_cast sigma c = match EConstr.kind sigma c with
| Cast (c,_,_) -> strip_outer_cast sigma c
diff --git a/engine/termops.mli b/engine/termops.mli
index ef2c52a45..2dab0685d 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -10,13 +10,13 @@
needed in the kernel. *)
open Names
-open Term
+open Constr
open Environ
open EConstr
(** printers *)
-val print_sort : sorts -> Pp.t
-val pr_sort_family : sorts_family -> Pp.t
+val print_sort : Sorts.t -> Pp.t
+val pr_sort_family : Sorts.family -> Pp.t
val pr_fix : ('a -> Pp.t) -> ('a, 'a) pfixpoint -> Pp.t
(** about contexts *)
@@ -147,7 +147,7 @@ val subst_term : Evd.evar_map -> constr -> constr -> constr
val replace_term : Evd.evar_map -> constr -> constr -> constr -> constr
(** Alternative term equalities *)
-val base_sort_cmp : Reduction.conv_pb -> sorts -> sorts -> bool
+val base_sort_cmp : Reduction.conv_pb -> Sorts.t -> Sorts.t -> bool
val compare_constr_univ : Evd.evar_map -> (Reduction.conv_pb -> constr -> constr -> bool) ->
Reduction.conv_pb -> constr -> constr -> bool
val constr_cmp : Evd.evar_map -> Reduction.conv_pb -> constr -> constr -> bool
diff --git a/engine/uState.ml b/engine/uState.ml
index 13a9bb373..77837fefc 100644
--- a/engine/uState.ml
+++ b/engine/uState.ml
@@ -33,10 +33,10 @@ type uinfo = {
(* 2nd part used to check consistency on the fly. *)
type t =
{ uctx_names : Univ.Level.t UNameMap.t * uinfo Univ.LMap.t;
- uctx_local : Univ.universe_context_set; (** The local context of variables *)
+ uctx_local : Univ.ContextSet.t; (** The local context of variables *)
uctx_univ_variables : Universes.universe_opt_subst;
(** The local universes that are unification variables *)
- uctx_univ_algebraic : Univ.universe_set;
+ uctx_univ_algebraic : Univ.LSet.t;
(** The subset of unification variables that can be instantiated with
algebraic universes as they appear in inferred types only. *)
uctx_universes : UGraph.t; (** The current graph extended with the local constraints *)
diff --git a/engine/uState.mli b/engine/uState.mli
index c44f2c1d7..b31e94b28 100644
--- a/engine/uState.mli
+++ b/engine/uState.mli
@@ -28,13 +28,13 @@ val is_empty : t -> bool
val union : t -> t -> t
-val of_context_set : Univ.universe_context_set -> t
+val of_context_set : Univ.ContextSet.t -> t
val of_binders : Universes.universe_binders -> t
(** {5 Projections} *)
-val context_set : t -> Univ.universe_context_set
+val context_set : t -> Univ.ContextSet.t
(** The local context of the state, i.e. a set of bound variables together
with their associated constraints. *)
@@ -54,7 +54,7 @@ val algebraics : t -> Univ.LSet.t
val constraints : t -> Univ.constraints
(** Shorthand for {!context_set} composed with {!ContextSet.constraints}. *)
-val context : t -> Univ.universe_context
+val context : t -> Univ.UContext.t
(** Shorthand for {!context_set} with {!Context_set.to_context}. *)
(** {5 Constraints handling} *)
@@ -79,7 +79,7 @@ val universe_of_name : t -> string -> Univ.Level.t
(** {5 Unification} *)
-val restrict : t -> Univ.universe_set -> t
+val restrict : t -> Univ.LSet.t -> t
type rigid =
| UnivRigid
@@ -89,7 +89,7 @@ val univ_rigid : rigid
val univ_flexible : rigid
val univ_flexible_alg : rigid
-val merge : ?loc:Loc.t -> bool -> rigid -> t -> Univ.universe_context_set -> t
+val merge : ?loc:Loc.t -> bool -> rigid -> t -> Univ.ContextSet.t -> t
val merge_subst : t -> Universes.universe_opt_subst -> t
val emit_side_effects : Safe_typing.private_constants -> t -> t
@@ -130,12 +130,12 @@ val normalize : t -> t
Also return the association list of universe names and universes
(including those not in [names]). *)
val universe_context : names:(Id.t Loc.located) list -> extensible:bool -> t ->
- (Id.t * Univ.Level.t) list * Univ.universe_context
+ (Id.t * Univ.Level.t) list * Univ.UContext.t
type universe_decl =
(Names.Id.t Loc.located list, Univ.Constraint.t) Misctypes.gen_universe_decl
-val check_univ_decl : t -> universe_decl -> Universes.universe_binders * Univ.universe_context
+val check_univ_decl : t -> universe_decl -> Universes.universe_binders * Univ.UContext.t
(** {5 TODO: Document me} *)
diff --git a/engine/universes.ml b/engine/universes.ml
index 0a2045a04..3136f805c 100644
--- a/engine/universes.ml
+++ b/engine/universes.ml
@@ -6,10 +6,11 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Sorts
open Util
open Pp
open Names
-open Term
+open Constr
open Environ
open Univ
open Globnames
@@ -20,7 +21,7 @@ let pr_with_global_universes l =
(** Local universe names of polymorphic references *)
-type universe_binders = (Id.t * Univ.universe_level) list
+type universe_binders = (Id.t * Univ.Level.t) list
let universe_binders_table = Summary.ref Refmap.empty ~name:"universe binders"
@@ -39,7 +40,7 @@ let is_set_minimization () = !set_minimization
type universe_constraint_type = ULe | UEq | ULub
-type universe_constraint = universe * universe_constraint_type * universe
+type universe_constraint = Universe.t * universe_constraint_type * Universe.t
module Constraints = struct
module S = Set.Make(
@@ -157,10 +158,10 @@ let eq_constr_univs_infer_with kind1 kind2 univs fold m n accu =
if res then Some !cstrs else None
let compare_head_gen_proj env equ eqs eqc' m n =
- match kind_of_term m, kind_of_term n with
+ match kind m, kind n with
| Proj (p, c), App (f, args)
| App (f, args), Proj (p, c) ->
- (match kind_of_term f with
+ (match kind f with
| Const (p', u) when Constant.equal (Projection.constant p) p' ->
let pb = Environ.lookup_projection p env in
let npars = pb.Declarations.proj_npars in
@@ -328,7 +329,7 @@ let fresh_global_or_constr_instance env = function
| IsGlobal gr -> fresh_global_instance env gr
let global_of_constr c =
- match kind_of_term c with
+ match kind c with
| Const (c, u) -> ConstRef c, u
| Ind (i, u) -> IndRef i, u
| Construct (c, u) -> ConstructRef c, u
@@ -390,8 +391,8 @@ let type_of_reference env r =
let type_of_global t = type_of_reference (Global.env ()) t
let fresh_sort_in_family env = function
- | InProp -> prop_sort, ContextSet.empty
- | InSet -> set_sort, ContextSet.empty
+ | InProp -> Sorts.prop, ContextSet.empty
+ | InSet -> Sorts.set, ContextSet.empty
| InType ->
let u = fresh_level () in
Type (Univ.Universe.make u), ContextSet.singleton u
@@ -449,7 +450,7 @@ let nf_evars_and_universes_opt_subst f subst =
let subst = fun l -> match LMap.find l subst with None -> raise Not_found | Some l' -> l' in
let lsubst = Univ.level_subst_of subst in
let rec aux c =
- match kind_of_term c with
+ match kind c with
| Evar (evk, args) ->
let args = Array.map aux args in
(match try f (evk, args) with Not_found -> None with
@@ -467,7 +468,7 @@ let nf_evars_and_universes_opt_subst f subst =
| Sort (Type u) ->
let u' = Univ.subst_univs_universe subst u in
if u' == u then c else mkSort (sort_of_univ u')
- | _ -> map_constr aux c
+ | _ -> Constr.map aux c
in aux
let fresh_universe_context_set_instance ctx =
@@ -526,7 +527,7 @@ let normalize_opt_subst ctx =
else try ignore(normalize u) with Not_found -> assert(false)) ctx
in !ectx
-type universe_opt_subst = universe option universe_map
+type universe_opt_subst = Universe.t option universe_map
let make_opt_subst s =
fun x ->
@@ -788,7 +789,7 @@ let normalize_context_set ctx us algs =
(* We first put constraints in a normal-form: all self-loops are collapsed
to equalities. *)
let g = Univ.LSet.fold (fun v g -> UGraph.add_universe v false g)
- ctx UGraph.empty_universes
+ ctx UGraph.initial_universes
in
let g =
Univ.Constraint.fold
diff --git a/engine/universes.mli b/engine/universes.mli
index 3678ec94d..24613c4b9 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -8,7 +8,7 @@
open Util
open Names
-open Term
+open Constr
open Environ
open Univ
@@ -21,26 +21,26 @@ val pr_with_global_universes : Level.t -> Pp.t
(** Local universe name <-> level mapping *)
-type universe_binders = (Id.t * Univ.universe_level) list
-
+type universe_binders = (Id.t * Univ.Level.t) list
+
val register_universe_binders : Globnames.global_reference -> universe_binders -> unit
val universe_binders_of_global : Globnames.global_reference -> universe_binders
(** The global universe counter *)
-val set_remote_new_univ_level : universe_level RemoteCounter.installer
+val set_remote_new_univ_level : Level.t RemoteCounter.installer
(** Side-effecting functions creating new universe levels. *)
-val new_univ_level : DirPath.t -> universe_level
-val new_univ : DirPath.t -> universe
+val new_univ_level : DirPath.t -> Level.t
+val new_univ : DirPath.t -> Universe.t
val new_Type : DirPath.t -> types
-val new_Type_sort : DirPath.t -> sorts
+val new_Type_sort : DirPath.t -> Sorts.t
-val new_global_univ : unit -> universe in_universe_context_set
-val new_sort_in_family : sorts_family -> sorts
+val new_global_univ : unit -> Universe.t in_universe_context_set
+val new_sort_in_family : Sorts.family -> Sorts.t
(** {6 Constraints for type inference}
-
+
When doing conversion of universes, not only do we have =/<= constraints but
also Lub constraints which correspond to unification of two levels which might
not be necessary if unfolding is performed.
@@ -48,7 +48,7 @@ val new_sort_in_family : sorts_family -> sorts
type universe_constraint_type = ULe | UEq | ULub
-type universe_constraint = universe * universe_constraint_type * universe
+type universe_constraint = Universe.t * universe_constraint_type * Universe.t
module Constraints : sig
include Set.S with type elt = universe_constraint
@@ -63,7 +63,7 @@ type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> unive
val subst_univs_universe_constraints : universe_subst_fn ->
universe_constraints -> universe_constraints
-val enforce_eq_instances_univs : bool -> universe_instance universe_constraint_function
+val enforce_eq_instances_univs : bool -> Instance.t universe_constraint_function
val to_constraints : UGraph.t -> universe_constraints -> constraints
@@ -82,14 +82,14 @@ val eq_constr_universes_proj : env -> constr -> constr -> bool universe_constrai
(** Build a fresh instance for a given context, its associated substitution and
the instantiated constraints. *)
-val fresh_instance_from_context : abstract_universe_context ->
- universe_instance constrained
+val fresh_instance_from_context : AUContext.t ->
+ Instance.t constrained
-val fresh_instance_from : abstract_universe_context -> universe_instance option ->
- universe_instance in_universe_context_set
+val fresh_instance_from : AUContext.t -> Instance.t option ->
+ Instance.t in_universe_context_set
-val fresh_sort_in_family : env -> sorts_family ->
- sorts in_universe_context_set
+val fresh_sort_in_family : env -> Sorts.family ->
+ Sorts.t in_universe_context_set
val fresh_constant_instance : env -> Constant.t ->
pconstant in_universe_context_set
val fresh_inductive_instance : env -> inductive ->
@@ -105,15 +105,15 @@ val fresh_global_or_constr_instance : env -> Globnames.global_reference_or_const
(** Get fresh variables for the universe context.
Useful to make tactics that manipulate constrs in universe contexts polymorphic. *)
-val fresh_universe_context_set_instance : universe_context_set ->
- universe_level_subst * universe_context_set
+val fresh_universe_context_set_instance : ContextSet.t ->
+ universe_level_subst * ContextSet.t
(** Raises [Not_found] if not a global reference. *)
val global_of_constr : constr -> Globnames.global_reference puniverses
val constr_of_global_univ : Globnames.global_reference puniverses -> constr
-val extend_context : 'a in_universe_context_set -> universe_context_set ->
+val extend_context : 'a in_universe_context_set -> ContextSet.t ->
'a in_universe_context_set
(** Simplification and pruning of constraints:
@@ -127,38 +127,38 @@ val extend_context : 'a in_universe_context_set -> universe_context_set ->
(a global one if there is one) and transitively saturate
the constraints w.r.t to the equalities. *)
-module UF : Unionfind.PartitionSig with type elt = universe_level
+module UF : Unionfind.PartitionSig with type elt = Level.t
-type universe_opt_subst = universe option universe_map
+type universe_opt_subst = Universe.t option universe_map
val make_opt_subst : universe_opt_subst -> universe_subst_fn
val subst_opt_univs_constr : universe_opt_subst -> constr -> constr
-val normalize_context_set : universe_context_set ->
+val normalize_context_set : ContextSet.t ->
universe_opt_subst (* The defined and undefined variables *) ->
- universe_set (* univ variables that can be substituted by algebraics *) ->
- (universe_opt_subst * universe_set) in_universe_context_set
+ LSet.t (* univ variables that can be substituted by algebraics *) ->
+ (universe_opt_subst * LSet.t) in_universe_context_set
val normalize_univ_variables : universe_opt_subst ->
- universe_opt_subst * universe_set * universe_set * universe_subst
+ universe_opt_subst * LSet.t * LSet.t * universe_subst
val normalize_univ_variable :
- find:(universe_level -> universe) ->
- update:(universe_level -> universe -> universe) ->
- universe_level -> universe
+ find:(Level.t -> Universe.t) ->
+ update:(Level.t -> Universe.t -> Universe.t) ->
+ Level.t -> Universe.t
val normalize_univ_variable_opt_subst : universe_opt_subst ref ->
- (universe_level -> universe)
+ (Level.t -> Universe.t)
val normalize_univ_variable_subst : universe_subst ref ->
- (universe_level -> universe)
+ (Level.t -> Universe.t)
val normalize_universe_opt_subst : universe_opt_subst ref ->
- (universe -> universe)
+ (Universe.t -> Universe.t)
val normalize_universe_subst : universe_subst ref ->
- (universe -> universe)
+ (Universe.t -> Universe.t)
(** Create a fresh global in the global environment, without side effects.
BEWARE: this raises an ANOMALY on polymorphic constants/inductives:
@@ -180,7 +180,7 @@ val type_of_global : Globnames.global_reference -> types in_universe_context_set
val nf_evars_and_universes_opt_subst : (existential -> constr option) ->
universe_opt_subst -> constr -> constr
-val refresh_constraints : UGraph.t -> universe_context_set -> universe_context_set * UGraph.t
+val refresh_constraints : UGraph.t -> ContextSet.t -> ContextSet.t * UGraph.t
(** Pretty-printing *)
@@ -188,11 +188,11 @@ val pr_universe_opt_subst : universe_opt_subst -> Pp.t
(** {6 Support for template polymorphism } *)
-val solve_constraints_system : universe option array -> universe array -> universe array ->
- universe array
+val solve_constraints_system : Universe.t option array -> Universe.t array -> Universe.t array ->
+ Universe.t array
(** Operations for universe_info_ind *)
(** Given a universe context representing constraints of an inductive
this function produces a UInfoInd.t that with the trivial subtyping relation. *)
-val univ_inf_ind_from_universe_context : universe_context -> cumulativity_info
+val univ_inf_ind_from_universe_context : UContext.t -> CumulativityInfo.t
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index bd6aa0911..e1cf8f196 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -12,7 +12,7 @@ open CErrors
open Util
open Names
open Nameops
-open Term
+open Constr
open Termops
open Libnames
open Globnames
diff --git a/interp/constrextern.mli b/interp/constrextern.mli
index b2df449c5..d980b1995 100644
--- a/interp/constrextern.mli
+++ b/interp/constrextern.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Names
-open Term
open Termops
open EConstr
open Environ
@@ -41,7 +40,7 @@ val extern_constr : ?lax:bool -> bool -> env -> Evd.evar_map -> constr -> constr
val extern_constr_in_scope : bool -> scope_name -> env -> Evd.evar_map -> constr -> constr_expr
val extern_reference : ?loc:Loc.t -> Id.Set.t -> global_reference -> reference
val extern_type : bool -> env -> Evd.evar_map -> types -> constr_expr
-val extern_sort : Evd.evar_map -> sorts -> glob_sort
+val extern_sort : Evd.evar_map -> Sorts.t -> glob_sort
val extern_rel_context : constr option -> env -> Evd.evar_map ->
rel_context -> local_binder_expr list
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index 75e99dd9b..46f96d20b 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Evd
open Environ
open Libnames
diff --git a/interp/declare.ml b/interp/declare.ml
index 0ead94eff..0cc4d0fca 100644
--- a/interp/declare.ml
+++ b/interp/declare.ml
@@ -15,7 +15,7 @@ open Names
open Libnames
open Globnames
open Nameops
-open Term
+open Constr
open Declarations
open Entries
open Libobject
@@ -440,7 +440,7 @@ let assumption_message id =
(** Global universe names, in a different summary *)
-type universe_context_decl = polymorphic * Univ.universe_context_set
+type universe_context_decl = polymorphic * Univ.ContextSet.t
let cache_universe_context (p, ctx) =
Global.push_context_set p ctx;
@@ -458,7 +458,7 @@ let declare_universe_context poly ctx =
Lib.add_anonymous_leaf (input_universe_context (poly, ctx))
(* Discharged or not *)
-type universe_decl = polymorphic * (Id.t * Univ.universe_level) list
+type universe_decl = polymorphic * (Id.t * Univ.Level.t) list
let cache_universes (p, l) =
let glob = Global.global_universe_names () in
diff --git a/interp/declare.mli b/interp/declare.mli
index 86d33cfb2..9b3194dec 100644
--- a/interp/declare.mli
+++ b/interp/declare.mli
@@ -8,7 +8,7 @@
open Names
open Libnames
-open Term
+open Constr
open Entries
open Decl_kinds
@@ -42,7 +42,7 @@ type internal_flag =
(* Defaut definition entries, transparent with no secctx or proj information *)
val definition_entry : ?fix_exn:Future.fix_exn ->
?opaque:bool -> ?inline:bool -> ?types:types ->
- ?poly:polymorphic -> ?univs:Univ.universe_context ->
+ ?poly:polymorphic -> ?univs:Univ.UContext.t ->
?eff:Safe_typing.private_constants -> constr -> Safe_typing.private_constants definition_entry
(** [declare_constant id cd] declares a global declaration
@@ -84,7 +84,7 @@ val exists_name : Id.t -> bool
(** Global universe contexts, names and constraints *)
-val declare_universe_context : polymorphic -> Univ.universe_context_set -> unit
+val declare_universe_context : polymorphic -> Univ.ContextSet.t -> unit
val do_universe : polymorphic -> Id.t Loc.located list -> unit
val do_constraint : polymorphic ->
diff --git a/interp/discharge.ml b/interp/discharge.ml
index 0e4bbd299..5b4b5f67b 100644
--- a/interp/discharge.ml
+++ b/interp/discharge.ml
@@ -10,6 +10,7 @@ open Names
open CErrors
open Util
open Term
+open Constr
open Vars
open Declarations
open Cooking
diff --git a/interp/impargs.ml b/interp/impargs.ml
index daec1191e..72d22db4d 100644
--- a/interp/impargs.ml
+++ b/interp/impargs.ml
@@ -12,6 +12,7 @@ open Names
open Globnames
open Nameops
open Term
+open Constr
open Reduction
open Declarations
open Environ
@@ -167,7 +168,7 @@ let update pos rig (na,st) =
(* modified is_rigid_reference with a truncated env *)
let is_flexible_reference env bound depth f =
- match kind_of_term f with
+ match kind f with
| Rel n when n >= bound+depth -> (* inductive type *) false
| Rel n when n >= depth -> (* previous argument *) true
| Rel n -> (* since local definitions have been expanded *) false
@@ -191,7 +192,7 @@ let add_free_rels_until strict strongly_strict revpat bound env m pos acc =
let rec frec rig (env,depth as ed) c =
let hd = if strict then whd_all env c else c in
let c = if strongly_strict then hd else c in
- match kind_of_term hd with
+ match kind hd with
| Rel n when (n < bound+depth) && (n >= depth) ->
let i = bound + depth - n - 1 in
acc.(i) <- update pos rig acc.(i)
@@ -214,13 +215,13 @@ let add_free_rels_until strict strongly_strict revpat bound env m pos acc =
let () = if not (Vars.noccur_between 1 bound m) then frec true (env,1) m in
acc
-let rec is_rigid_head t = match kind_of_term t with
+let rec is_rigid_head t = match kind t with
| Rel _ | Evar _ -> false
| Ind _ | Const _ | Var _ | Sort _ -> true
| Case (_,_,f,_) -> is_rigid_head f
| Proj (p,c) -> true
| App (f,args) ->
- (match kind_of_term f with
+ (match kind f with
| Fix ((fi,i),_) -> is_rigid_head (args.(fi.(i)))
| _ -> is_rigid_head f)
| Lambda _ | LetIn _ | Construct _ | CoFix _ | Fix _
@@ -240,7 +241,7 @@ let compute_implicits_gen strict strongly_strict revpat contextual all env t =
let open Context.Rel.Declaration in
let rec aux env avoid n names t =
let t = whd_all env t in
- match kind_of_term t with
+ match kind t with
| Prod (na,a,b) ->
let na',avoid' = find_displayed_name_in all avoid na (names,b) in
add_free_rels_until strict strongly_strict revpat n env a (Hyp (n+1))
@@ -253,7 +254,7 @@ let compute_implicits_gen strict strongly_strict revpat contextual all env t =
add_free_rels_until strict strongly_strict revpat n env t Conclusion v
else v
in
- match kind_of_term (whd_all env t) with
+ match kind (whd_all env t) with
| Prod (na,a,b) ->
let na',avoid = find_displayed_name_in all Id.Set.empty na ([],b) in
let v = aux (push_rel (LocalAssum (na',a)) env) avoid 1 [na'] b in
diff --git a/interp/impargs.mli b/interp/impargs.mli
index 658d9e01a..40fa4cb26 100644
--- a/interp/impargs.mli
+++ b/interp/impargs.mli
@@ -7,8 +7,8 @@
(************************************************************************)
open Names
+open Constr
open Globnames
-open Term
open Environ
(** {6 Implicit Arguments } *)
diff --git a/interp/notation.ml b/interp/notation.ml
index 076570c8a..f36294f73 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -11,7 +11,7 @@ open CErrors
open Util
open Pp
open Names
-open Term
+open Constr
open Libnames
open Globnames
open Constrexpr
@@ -653,7 +653,7 @@ let find_scope_class_opt = function
(* Special scopes associated to arguments of a global reference *)
let rec compute_arguments_classes t =
- match kind_of_term (EConstr.Unsafe.to_constr (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t))) with
+ match Constr.kind (EConstr.Unsafe.to_constr (Reductionops.whd_betaiotazeta Evd.empty (EConstr.of_constr t))) with
| Prod (_,t,u) ->
let cl = try Some (compute_scope_class t) with Not_found -> None in
cl :: compute_arguments_classes u
diff --git a/interp/notation.mli b/interp/notation.mli
index e312a3767..2066d346f 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -165,8 +165,8 @@ val subst_scope_class :
val declare_scope_class : scope_name -> scope_class -> unit
val declare_ref_arguments_scope : global_reference -> unit
-val compute_arguments_scope : Term.types -> scope_name option list
-val compute_type_scope : Term.types -> scope_name option
+val compute_arguments_scope : Constr.types -> scope_name option list
+val compute_type_scope : Constr.types -> scope_name option
(** Get the current scope bound to Sortclass, if it exists *)
val current_type_scope_name : unit -> scope_name option
diff --git a/intf/constrexpr.ml b/intf/constrexpr.ml
index 8eadafe66..e0d2d7bf4 100644
--- a/intf/constrexpr.ml
+++ b/intf/constrexpr.ml
@@ -79,7 +79,7 @@ and constr_expr_r =
| CRecord of (reference * constr_expr) list
(* representation of the "let" and "match" constructs *)
- | CCases of case_style (* determines whether this value represents "let" or "match" construct *)
+ | CCases of Constr.case_style (* determines whether this value represents "let" or "match" construct *)
* constr_expr option (* return-clause *)
* case_expr list
* branch_expr list (* branches *)
diff --git a/intf/glob_term.ml b/intf/glob_term.ml
index 508990a58..72c91db6a 100644
--- a/intf/glob_term.ml
+++ b/intf/glob_term.ml
@@ -46,7 +46,7 @@ type 'a glob_constr_r =
| GLambda of Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
| GProd of Name.t * binding_kind * 'a glob_constr_g * 'a glob_constr_g
| GLetIn of Name.t * 'a glob_constr_g * 'a glob_constr_g option * 'a glob_constr_g
- | GCases of case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
+ | GCases of Constr.case_style * 'a glob_constr_g option * 'a tomatch_tuples_g * 'a cases_clauses_g
(** [GCases(style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in [MatchStyle]) *)
| GLetTuple of Name.t list * (Name.t * 'a glob_constr_g option) * 'a glob_constr_g * 'a glob_constr_g
| GIf of 'a glob_constr_g * (Name.t * 'a glob_constr_g option) * 'a glob_constr_g * 'a glob_constr_g
diff --git a/intf/misctypes.ml b/intf/misctypes.ml
index 8b7073143..87484ccd5 100644
--- a/intf/misctypes.ml
+++ b/intf/misctypes.ml
@@ -61,12 +61,13 @@ type existential_key = Evar.t
(** Case style, shared with Term *)
-type case_style = Term.case_style =
+type case_style = Constr.case_style =
| LetStyle
| IfStyle
| LetPatternStyle
| MatchStyle
| RegularStyle (** infer printing form from number of constructor *)
+[@@ocaml.deprecated "Alias for Constr.case_style"]
(** Casts *)
diff --git a/intf/notation_term.ml b/intf/notation_term.ml
index c342da3dc..7823d3feb 100644
--- a/intf/notation_term.ml
+++ b/intf/notation_term.ml
@@ -31,7 +31,7 @@ type notation_constr =
| NProd of Name.t * notation_constr * notation_constr
| NBinderList of Id.t * Id.t * notation_constr * notation_constr
| NLetIn of Name.t * notation_constr * notation_constr option * notation_constr
- | NCases of case_style * notation_constr option *
+ | NCases of Constr.case_style * notation_constr option *
(notation_constr * (Name.t * (inductive * Name.t list) option)) list *
(cases_pattern list * notation_constr) list
| NLetTuple of Name.t list * (Name.t * notation_constr option) *
diff --git a/intf/pattern.ml b/intf/pattern.ml
index 16c480735..20636accf 100644
--- a/intf/pattern.ml
+++ b/intf/pattern.ml
@@ -8,13 +8,13 @@
open Names
open Globnames
-open Term
+open Constr
open Misctypes
(** {5 Patterns} *)
type case_info_pattern =
- { cip_style : case_style;
+ { cip_style : Constr.case_style;
cip_ind : inductive option;
cip_ind_tags : bool list option; (** indicates LetIn/Lambda in arity *)
cip_extensible : bool (** does this match end with _ => _ ? *) }
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index e8cce0841..e1b086b75 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -23,7 +23,7 @@ open CErrors
open Util
open Pp
open Names
-open Term
+open Constr
open Vars
open Environ
open Esubst
@@ -516,7 +516,7 @@ let zupdate m s =
else s
let mk_lambda env t =
- let (rvars,t') = decompose_lam t in
+ let (rvars,t') = Term.decompose_lam t in
FLambda(List.length rvars, List.rev rvars, t', env)
let destFLambda clos_fun t =
@@ -530,7 +530,7 @@ let destFLambda clos_fun t =
(* Optimization: do not enclose variables in a closure.
Makes variable access much faster *)
let mk_clos e t =
- match kind_of_term t with
+ match kind t with
| Rel i -> clos_rel e i
| Var x -> { norm = Red; term = FFlex (VarKey x) }
| Const c -> { norm = Red; term = FFlex (ConstKey c) }
@@ -556,7 +556,7 @@ let mk_clos_vect env v = match v with
subterms.
Could be used insted of mk_clos. *)
let mk_clos_deep clos_fun env t =
- match kind_of_term t with
+ match kind t with
| (Rel _|Ind _|Const _|Construct _|Var _|Meta _ | Sort _) ->
mk_clos env t
| Cast (a,k,b) ->
@@ -655,7 +655,7 @@ let term_of_fconstr =
match v.term with
| FCLOS(t,env) when is_subs_id env && is_lift_id lfts -> t
| FLambda(_,tys,f,e) when is_subs_id e && is_lift_id lfts ->
- compose_lam (List.rev tys) f
+ Term.compose_lam (List.rev tys) f
| FFix(fx,e) when is_subs_id e && is_lift_id lfts -> mkFix fx
| FCoFix(cfx,e) when is_subs_id e && is_lift_id lfts -> mkCoFix cfx
| _ -> to_constr term_of_fconstr_lift lfts v in
@@ -891,7 +891,7 @@ let rec knh info m stk =
(* The same for pure terms *)
and knht info e t stk =
- match kind_of_term t with
+ match kind t with
| App(a,b) ->
knht info e a (append_stack (mk_clos_vect e b) stk)
| Case(ci,p,t,br) ->
diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli
index de1bb120e..28136e1fc 100644
--- a/kernel/cClosure.mli
+++ b/kernel/cClosure.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Esubst
diff --git a/kernel/cbytecodes.ml b/kernel/cbytecodes.ml
index 63b0e6929..9febc6449 100644
--- a/kernel/cbytecodes.ml
+++ b/kernel/cbytecodes.ml
@@ -13,7 +13,7 @@
(* This file defines the type of bytecode instructions *)
open Names
-open Term
+open Constr
type tag = int
@@ -32,13 +32,13 @@ let cofix_evaluated_tag = 7
let last_variant_tag = 245
type structured_constant =
- | Const_sorts of sorts
+ | Const_sorts of Sorts.t
| Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
- | Const_univ_level of Univ.universe_level
- | Const_type of Univ.universe
+ | Const_univ_level of Univ.Level.t
+ | Const_type of Univ.Universe.t
type reloc_table = (tag * int) array
@@ -186,7 +186,8 @@ open Pp
open Util
let pp_sort s =
- match family_of_sort s with
+ let open Sorts in
+ match family s with
| InSet -> str "Set"
| InProp -> str "Prop"
| InType -> str "Type"
diff --git a/kernel/cbytecodes.mli b/kernel/cbytecodes.mli
index e9070b26a..5d37a5840 100644
--- a/kernel/cbytecodes.mli
+++ b/kernel/cbytecodes.mli
@@ -9,7 +9,7 @@
(* $Id$ *)
open Names
-open Term
+open Constr
type tag = int
@@ -26,13 +26,13 @@ val cofix_evaluated_tag : tag
val last_variant_tag : tag
type structured_constant =
- | Const_sorts of sorts
+ | Const_sorts of Sorts.t
| Const_ind of inductive
| Const_proj of Constant.t
| Const_b0 of tag
| Const_bn of tag * structured_constant array
- | Const_univ_level of Univ.universe_level
- | Const_type of Univ.universe
+ | Const_univ_level of Univ.Level.t
+ | Const_type of Univ.Universe.t
val pp_struct_const : structured_constant -> Pp.t
diff --git a/kernel/cbytegen.ml b/kernel/cbytegen.ml
index 5bac26bf9..5dab2932d 100644
--- a/kernel/cbytegen.ml
+++ b/kernel/cbytegen.ml
@@ -14,7 +14,7 @@ open Util
open Names
open Cbytecodes
open Cemitcodes
-open Term
+open Constr
open Declarations
open Pre_env
@@ -438,12 +438,12 @@ let get_strcst = function
| _ -> raise Not_found
let rec str_const c =
- match kind_of_term c with
+ match kind c with
| Sort s -> Bstrconst (Const_sorts s)
| Cast(c,_,_) -> str_const c
| App(f,args) ->
begin
- match kind_of_term f with
+ match kind f with
| Construct(((kn,j),i),u) ->
begin
let oib = lookup_mind kn !global_env in
@@ -596,7 +596,7 @@ let rec get_alias env kn =
(* sz is the size of the local stack *)
let rec compile_constr reloc c sz cont =
set_max_stack_size sz;
- match kind_of_term c with
+ match kind c with
| Meta _ -> invalid_arg "Cbytegen.compile_constr : Meta"
| Evar _ -> invalid_arg "Cbytegen.compile_constr : Evar"
| Proj (p,c) ->
@@ -621,9 +621,9 @@ let rec compile_constr reloc c sz cont =
(Univ.Instance.to_array u)
sz
cont
- | Sort (Prop _) | Construct _ ->
+ | Sort (Sorts.Prop _) | Construct _ ->
compile_str_cst reloc (str_const c) sz cont
- | Sort (Type u) ->
+ | Sort (Sorts.Type u) ->
(* We separate global and local universes in [u]. The former will be part
of the structured constant, while the later (if any) will be applied as
arguments. *)
@@ -641,7 +641,7 @@ let rec compile_constr reloc c sz cont =
LSet.fold (fun lvl u -> Universe.sup u (Universe.make lvl)) global_levels Universe.type0m
in
if local_levels = [] then
- compile_str_cst reloc (Bstrconst (Const_sorts (Type uglob))) sz cont
+ compile_str_cst reloc (Bstrconst (Const_sorts (Sorts.Type uglob))) sz cont
else
let compile_get_univ reloc idx sz cont =
set_max_stack_size sz;
@@ -659,7 +659,7 @@ let rec compile_constr reloc c sz cont =
Kpush :: compile_constr reloc dom (sz+1) (Kmakeprod :: cont) in
compile_constr reloc (mkLambda(id,dom,codom)) sz cont1
| Lambda _ ->
- let params, body = decompose_lam c in
+ let params, body = Term.decompose_lam c in
let arity = List.length params in
let r_fun = comp_env_fun arity in
let lbl_fun = Label.create() in
@@ -672,7 +672,7 @@ let rec compile_constr reloc c sz cont =
| App(f,args) ->
begin
- match kind_of_term f with
+ match kind f with
| Construct _ -> compile_str_cst reloc (str_const c) sz cont
| Const (kn,u) -> compile_const reloc kn u args sz cont
| _ -> comp_app compile_constr compile_constr reloc f args sz cont
@@ -694,7 +694,7 @@ let rec compile_constr reloc c sz cont =
done;
(* Compiling bodies *)
for i = 0 to ndef - 1 do
- let params,body = decompose_lam rec_bodies.(i) in
+ let params,body = Term.decompose_lam rec_bodies.(i) in
let arity = List.length params in
let env_body = comp_env_fix ndef i arity rfv in
let cont1 =
@@ -726,7 +726,7 @@ let rec compile_constr reloc c sz cont =
done;
(* Compiling bodies *)
for i = 0 to ndef - 1 do
- let params,body = decompose_lam rec_bodies.(i) in
+ let params,body = Term.decompose_lam rec_bodies.(i) in
let arity = List.length params in
let env_body = comp_env_cofix ndef arity rfv in
let lbl = Label.create () in
@@ -792,7 +792,7 @@ let rec compile_constr reloc c sz cont =
lbl_consts.(tag) <- lbl_b;
c := code_b
else
- let args, body = decompose_lam branchs.(i) in
+ let args, body = Term.decompose_lam branchs.(i) in
let nargs = List.length args in
let code_b =
@@ -953,9 +953,9 @@ let compile fail_on_error ?universes:(universes=0) env c =
*)
let reloc = empty_comp_env () in
let arity , body =
- match kind_of_term c with
+ match kind c with
| Lambda _ ->
- let params, body = decompose_lam c in
+ let params, body = Term.decompose_lam c in
List.length params , body
| _ -> 0 , c
in
@@ -995,7 +995,7 @@ let compile_constant_body fail_on_error env univs = function
| Monomorphic_const _ -> 0
| Polymorphic_const univ -> Univ.AUContext.size univ
in
- match kind_of_term body with
+ match kind body with
| Const (kn',u) when is_univ_copy instance_size u ->
(* we use the canonical name of the constant*)
let con= Constant.make1 (Constant.canonical kn') in
@@ -1024,7 +1024,7 @@ let compile_structured_int31 fc args =
if not fc then raise Not_found else
Const_b0
(Array.fold_left
- (fun temp_i -> fun t -> match kind_of_term t with
+ (fun temp_i -> fun t -> match kind t with
| Construct ((_,d),_) -> 2*temp_i+d-1
| _ -> raise NotClosed)
0 args
diff --git a/kernel/cbytegen.mli b/kernel/cbytegen.mli
index 9209e8460..c117d3fb5 100644
--- a/kernel/cbytegen.mli
+++ b/kernel/cbytegen.mli
@@ -1,6 +1,14 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
open Cbytecodes
open Cemitcodes
-open Term
+open Constr
open Declarations
open Pre_env
diff --git a/kernel/constr.mli b/kernel/constr.mli
index da074d85c..474ab3884 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -87,7 +87,7 @@ val mkEvar : existential -> constr
val mkSort : Sorts.t -> types
val mkProp : types
val mkSet : types
-val mkType : Univ.universe -> types
+val mkType : Univ.Universe.t -> types
(** This defines the strategy to use for verifiying a Cast *)
diff --git a/kernel/cooking.ml b/kernel/cooking.ml
index 08fff0ca4..2579ac045 100644
--- a/kernel/cooking.ml
+++ b/kernel/cooking.ml
@@ -17,6 +17,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Declarations
open Univ
@@ -100,42 +101,42 @@ let expmod_constr cache modlist c =
let share_univs = share_univs cache in
let update_case_info = update_case_info cache in
let rec substrec c =
- match kind_of_term c with
+ match kind c with
| Case (ci,p,t,br) ->
- map_constr substrec (mkCase (update_case_info ci modlist,p,t,br))
+ Constr.map substrec (mkCase (update_case_info ci modlist,p,t,br))
| Ind (ind,u) ->
(try
share_univs (IndRef ind) u modlist
with
- | Not_found -> map_constr substrec c)
+ | Not_found -> Constr.map substrec c)
| Construct (cstr,u) ->
(try
share_univs (ConstructRef cstr) u modlist
with
- | Not_found -> map_constr substrec c)
+ | Not_found -> Constr.map substrec c)
| Const (cst,u) ->
(try
share_univs (ConstRef cst) u modlist
with
- | Not_found -> map_constr substrec c)
+ | Not_found -> Constr.map substrec c)
| Proj (p, c') ->
(try
let p' = share_univs (ConstRef (Projection.constant p)) Univ.Instance.empty modlist in
let make c = Projection.make c (Projection.unfolded p) in
- match kind_of_term p' with
+ match kind p' with
| Const (p',_) -> mkProj (make p', substrec c')
| App (f, args) ->
- (match kind_of_term f with
+ (match kind f with
| Const (p', _) -> mkProj (make p', substrec c')
| _ -> assert false)
| _ -> assert false
- with Not_found -> map_constr substrec c)
+ with Not_found -> Constr.map substrec c)
- | _ -> map_constr substrec c
+ | _ -> Constr.map substrec c
in
if is_empty_modlist modlist then c
@@ -203,7 +204,7 @@ let cook_constant ~hcons env { from = cb; info } =
let hyps = Context.Named.map expmod abstract in
let map c =
let c = abstract_constant_body (expmod c) hyps in
- if hcons then hcons_constr c else c
+ if hcons then Constr.hcons c else c
in
let body = on_body modlist (hyps, usubst, abs_ctx)
map
@@ -222,7 +223,7 @@ let cook_constant ~hcons env { from = cb; info } =
let ((mind, _), _), n' =
try
let c' = share_univs cache (IndRef (pb.proj_ind,0)) Univ.Instance.empty modlist in
- match kind_of_term c' with
+ match kind c' with
| App (f,l) -> (destInd f, Array.length l)
| Ind ind -> ind, 0
| _ -> assert false
diff --git a/kernel/cooking.mli b/kernel/cooking.mli
index 6d1b776c0..7696d7545 100644
--- a/kernel/cooking.mli
+++ b/kernel/cooking.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Declarations
open Environ
@@ -26,7 +26,7 @@ type result = {
}
val cook_constant : hcons:bool -> env -> recipe -> result
-val cook_constr : Opaqueproof.cooking_info -> Term.constr -> Term.constr
+val cook_constr : Opaqueproof.cooking_info -> constr -> constr
(** {6 Utility functions used in module [Discharge]. } *)
diff --git a/kernel/csymtable.ml b/kernel/csymtable.ml
index d21ea9670..2ffe36fcf 100644
--- a/kernel/csymtable.ml
+++ b/kernel/csymtable.ml
@@ -14,7 +14,7 @@
open Util
open Names
-open Term
+open Constr
open Vm
open Cemitcodes
open Cbytecodes
@@ -68,7 +68,7 @@ let rec eq_structured_constant c1 c2 = match c1, c2 with
| Const_bn (t1, a1), Const_bn (t2, a2) ->
Int.equal t1 t2 && Array.equal eq_structured_constant a1 a2
| Const_bn _, _ -> false
-| Const_univ_level l1 , Const_univ_level l2 -> Univ.eq_levels l1 l2
+| Const_univ_level l1 , Const_univ_level l2 -> Univ.Level.equal l1 l2
| Const_univ_level _ , _ -> false
| Const_type u1 , Const_type u2 -> Univ.Universe.equal u1 u2
| Const_type _ , _ -> false
diff --git a/kernel/csymtable.mli b/kernel/csymtable.mli
index 30d0e5588..91bb30e7e 100644
--- a/kernel/csymtable.mli
+++ b/kernel/csymtable.mli
@@ -9,7 +9,7 @@
(* $Id$ *)
open Names
-open Term
+open Constr
open Pre_env
val val_of_constr : env -> constr -> values
diff --git a/kernel/declarations.ml b/kernel/declarations.ml
index 8b949f928..b95796fd8 100644
--- a/kernel/declarations.ml
+++ b/kernel/declarations.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
(** This module defines the internal representation of global
declarations. This includes global constants/axioms, mutual
@@ -28,8 +28,8 @@ type engagement = set_predicativity
*)
type template_arity = {
- template_param_levels : Univ.universe_level option list;
- template_level : Univ.universe;
+ template_param_levels : Univ.Level.t option list;
+ template_level : Univ.Universe.t;
}
type ('a, 'b) declaration_arity =
@@ -63,8 +63,8 @@ type constant_def =
| OpaqueDef of Opaqueproof.opaque (** or an opaque global definition *)
type constant_universes =
- | Monomorphic_const of Univ.universe_context
- | Polymorphic_const of Univ.abstract_universe_context
+ | Monomorphic_const of Univ.UContext.t
+ | Polymorphic_const of Univ.AUContext.t
(** The [typing_flags] are instructions to the type-checker which
modify its behaviour. The typing flags used in the type-checking
@@ -119,7 +119,7 @@ type record_body = (Id.t * Constant.t array * projection_body array) option
type regular_inductive_arity = {
mind_user_arity : types;
- mind_sort : sorts;
+ mind_sort : Sorts.t;
}
type inductive_arity = (regular_inductive_arity, template_arity) declaration_arity
@@ -146,7 +146,7 @@ type one_inductive_body = {
mind_nrealdecls : int; (** Length of realargs context (with let, no params) *)
- mind_kelim : sorts_family list; (** List of allowed elimination sorts *)
+ mind_kelim : Sorts.family list; (** List of allowed elimination sorts *)
mind_nf_lc : types array; (** Head normalized constructor types so that their conclusion exposes the inductive type *)
@@ -168,9 +168,9 @@ type one_inductive_body = {
}
type abstract_inductive_universes =
- | Monomorphic_ind of Univ.universe_context
- | Polymorphic_ind of Univ.abstract_universe_context
- | Cumulative_ind of Univ.abstract_cumulativity_info
+ | Monomorphic_ind of Univ.UContext.t
+ | Polymorphic_ind of Univ.AUContext.t
+ | Cumulative_ind of Univ.ACumulativityInfo.t
type mutual_inductive_body = {
diff --git a/kernel/declareops.ml b/kernel/declareops.ml
index d7329a319..f5c26b33d 100644
--- a/kernel/declareops.ml
+++ b/kernel/declareops.ml
@@ -112,7 +112,7 @@ let subst_const_body sub cb =
themselves. But would it really bring substantial gains ? *)
let hcons_rel_decl =
- RelDecl.map_name Names.Name.hcons %> RelDecl.map_value Term.hcons_constr %> RelDecl.map_type Term.hcons_types
+ RelDecl.map_name Names.Name.hcons %> RelDecl.map_value Constr.hcons %> RelDecl.map_type Constr.hcons
let hcons_rel_context l = List.smartmap hcons_rel_decl l
@@ -120,7 +120,7 @@ let hcons_const_def = function
| Undef inl -> Undef inl
| Def l_constr ->
let constr = force_constr l_constr in
- Def (from_val (Term.hcons_constr constr))
+ Def (from_val (Constr.hcons constr))
| OpaqueDef _ as x -> x (* hashconsed when turned indirect *)
let hcons_const_universes cbu =
@@ -133,7 +133,7 @@ let hcons_const_universes cbu =
let hcons_const_body cb =
{ cb with
const_body = hcons_const_def cb.const_body;
- const_type = Term.hcons_constr cb.const_type;
+ const_type = Constr.hcons cb.const_type;
const_universes = hcons_const_universes cb.const_universes }
(** {6 Inductive types } *)
@@ -249,8 +249,8 @@ let inductive_is_cumulative mib =
(** {6 Hash-consing of inductive declarations } *)
let hcons_regular_ind_arity a =
- { mind_user_arity = Term.hcons_constr a.mind_user_arity;
- mind_sort = Term.hcons_sorts a.mind_sort }
+ { mind_user_arity = Constr.hcons a.mind_user_arity;
+ mind_sort = Sorts.hcons a.mind_sort }
(** Just as for constants, this hash-consing is quite partial *)
@@ -260,8 +260,8 @@ let hcons_ind_arity =
(** Substitution of inductive declarations *)
let hcons_mind_packet oib =
- let user = Array.smartmap Term.hcons_types oib.mind_user_lc in
- let nf = Array.smartmap Term.hcons_types oib.mind_nf_lc in
+ let user = Array.smartmap Constr.hcons oib.mind_user_lc in
+ let nf = Array.smartmap Constr.hcons oib.mind_nf_lc in
(* Special optim : merge [mind_user_lc] and [mind_nf_lc] if possible *)
let nf = if Array.equal (==) user nf then user else nf in
{ oib with
diff --git a/kernel/declareops.mli b/kernel/declareops.mli
index b2d29759d..198831848 100644
--- a/kernel/declareops.mli
+++ b/kernel/declareops.mli
@@ -27,7 +27,7 @@ val subst_const_body : substitution -> constant_body -> constant_body
val constant_has_body : constant_body -> bool
-val constant_polymorphic_context : constant_body -> abstract_universe_context
+val constant_polymorphic_context : constant_body -> AUContext.t
(** Is the constant polymorphic? *)
val constant_is_polymorphic : constant_body -> bool
@@ -57,7 +57,7 @@ val subst_wf_paths : substitution -> wf_paths -> wf_paths
val subst_mind_body : substitution -> mutual_inductive_body -> mutual_inductive_body
-val inductive_polymorphic_context : mutual_inductive_body -> abstract_universe_context
+val inductive_polymorphic_context : mutual_inductive_body -> AUContext.t
(** Is the inductive polymorphic? *)
val inductive_is_polymorphic : mutual_inductive_body -> bool
diff --git a/kernel/entries.ml b/kernel/entries.ml
index f294c4ce4..185dba409 100644
--- a/kernel/entries.ml
+++ b/kernel/entries.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
(** This module defines the entry types for global declarations. This
information is entered in the environments. This includes global
@@ -35,9 +35,9 @@ then, in i{^ th} block, [mind_entry_params] is [xn:Xn;...;x1:X1];
*)
type inductive_universes =
- | Monomorphic_ind_entry of Univ.universe_context
- | Polymorphic_ind_entry of Univ.universe_context
- | Cumulative_ind_entry of Univ.cumulativity_info
+ | Monomorphic_ind_entry of Univ.UContext.t
+ | Polymorphic_ind_entry of Univ.UContext.t
+ | Cumulative_ind_entry of Univ.CumulativityInfo.t
type one_inductive_entry = {
mind_entry_typename : Id.t;
@@ -65,8 +65,8 @@ type 'a proof_output = constr Univ.in_universe_context_set * 'a
type 'a const_entry_body = 'a proof_output Future.computation
type constant_universes_entry =
- | Monomorphic_const_entry of Univ.universe_context
- | Polymorphic_const_entry of Univ.universe_context
+ | Monomorphic_const_entry of Univ.UContext.t
+ | Polymorphic_const_entry of Univ.UContext.t
type 'a definition_entry = {
const_entry_body : 'a const_entry_body;
@@ -112,7 +112,7 @@ type seff_env =
[ `Nothing
(* The proof term and its universes.
Same as the constant_body's but not in an ephemeron *)
- | `Opaque of Constr.t * Univ.universe_context_set ]
+ | `Opaque of Constr.t * Univ.ContextSet.t ]
type side_eff =
| SEsubproof of Constant.t * Declarations.constant_body * seff_env
diff --git a/kernel/environ.ml b/kernel/environ.ml
index c3fd8962e..1afab453a 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -23,7 +23,7 @@
open CErrors
open Util
open Names
-open Term
+open Constr
open Vars
open Declarations
open Pre_env
@@ -391,7 +391,7 @@ let lookup_constructor_variables (ind,_) env =
(* Returns the list of global variables in a term *)
let vars_of_global env constr =
- match kind_of_term constr with
+ match kind constr with
Var id -> Id.Set.singleton id
| Const (kn, _) -> lookup_constant_variables kn env
| Ind (ind, _) -> lookup_inductive_variables ind env
@@ -402,12 +402,12 @@ let vars_of_global env constr =
let global_vars_set env constr =
let rec filtrec acc c =
let acc =
- match kind_of_term c with
+ match kind c with
| Var _ | Const _ | Ind _ | Construct _ ->
Id.Set.union (vars_of_global env c) acc
| _ ->
acc in
- Term.fold_constr filtrec acc c
+ Constr.fold filtrec acc c
in
filtrec Id.Set.empty constr
@@ -478,7 +478,7 @@ let j_type j = j.uj_type
type 'types punsafe_type_judgment = {
utj_val : 'types;
- utj_type : sorts }
+ utj_type : Sorts.t }
type unsafe_type_judgment = types punsafe_type_judgment
@@ -538,7 +538,7 @@ let register_one env field entry =
let register env field entry =
match field with
| KInt31 (grp, Int31Type) ->
- let i31c = match kind_of_term entry with
+ let i31c = match kind entry with
| Ind i31t -> mkConstructUi (i31t, 1)
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be an inductive type.")
in
@@ -584,7 +584,7 @@ let dispatch =
fun rk value field ->
(* subfunction which shortens the (very common) dispatch of operations *)
let int31_op_from_const n op prim =
- match kind_of_term value with
+ match kind value with
| Const kn -> int31_op n op prim kn
| _ -> anomaly ~label:"Environ.register" (Pp.str "should be a constant.")
in
@@ -601,13 +601,13 @@ fun rk value field ->
(Pp.str "add_int31_decompilation_from_type called with an abnormal field.")
in
let i31bit_type =
- match kind_of_term int31bit with
+ match kind int31bit with
| Ind (i31bit_type,_) -> i31bit_type
| _ -> anomaly ~label:"Environ.register"
(Pp.str "Int31Bits should be an inductive type.")
in
let int31_decompilation =
- match kind_of_term value with
+ match kind value with
| Ind (i31t,_) ->
constr_of_int31 i31t i31bit_type
| _ -> anomaly ~label:"Environ.register"
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 37816f1e5..f2066b065 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -7,9 +7,9 @@
(************************************************************************)
open Names
-open Term
-open Declarations
open Univ
+open Constr
+open Declarations
(** Unsafe environments. We define here a datatype for environments.
Since typing is not yet defined, it is not possible to check the
@@ -157,7 +157,7 @@ val constant_value_and_type : env -> Constant.t puniverses ->
constr option * types * Univ.constraints
(** The universe context associated to the constant, empty if not
polymorphic *)
-val constant_context : env -> Constant.t -> Univ.abstract_universe_context
+val constant_context : env -> Constant.t -> Univ.AUContext.t
(* These functions should be called under the invariant that [env]
already contains the constraints corresponding to the constant
@@ -207,8 +207,8 @@ val add_constraints : Univ.constraints -> env -> env
(** Check constraints are satifiable in the environment. *)
val check_constraints : Univ.constraints -> env -> bool
-val push_context : ?strict:bool -> Univ.universe_context -> env -> env
-val push_context_set : ?strict:bool -> Univ.universe_context_set -> env -> env
+val push_context : ?strict:bool -> Univ.UContext.t -> env -> env
+val push_context_set : ?strict:bool -> Univ.ContextSet.t -> env -> env
val push_constraints_to_env : 'a Univ.constrained -> env -> env
val set_engagement : engagement -> env -> env
@@ -247,7 +247,7 @@ val j_type : ('constr, 'types) punsafe_judgment -> 'types
type 'types punsafe_type_judgment = {
utj_val : 'types;
- utj_type : sorts }
+ utj_type : Sorts.t }
type unsafe_type_judgment = types punsafe_type_judgment
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index c0f564dc3..f4e611c19 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -11,6 +11,7 @@ open Util
open Names
open Univ
open Term
+open Constr
open Vars
open Declarations
open Declareops
@@ -55,7 +56,7 @@ let weaker_noccur_between env x nvars t =
else None
let is_constructor_head t =
- isRel(fst(decompose_app t))
+ Term.isRel(fst(Term.decompose_app t))
(************************************************************************)
(* Various well-formedness check for inductive declarations *)
@@ -130,11 +131,11 @@ let is_unit constrsinfos =
let infos_and_sort env t =
let rec aux env t max =
let t = whd_all env t in
- match kind_of_term t with
+ match kind t with
| Prod (name,c1,c2) ->
let varj = infer_type env c1 in
let env1 = Environ.push_rel (LocalAssum (name,varj.utj_val)) env in
- let max = Universe.sup max (univ_of_sort varj.utj_type) in
+ let max = Universe.sup max (Term.univ_of_sort varj.utj_type) in
aux env1 c2 max
| _ when is_constructor_head t -> max
| _ -> (* don't fail if not positive, it is tested later *) max
@@ -168,7 +169,7 @@ let infer_constructor_packet env_ar_par params lc =
let jlc = List.map (infer_type env_ar_par) lc in
let jlc = Array.of_list jlc in
(* generalize the constructor over the parameters *)
- let lc'' = Array.map (fun j -> it_mkProd_or_LetIn j.utj_val params) jlc in
+ let lc'' = Array.map (fun j -> Term.it_mkProd_or_LetIn j.utj_val params) jlc in
(* compute the max of the sorts of the products of the constructors types *)
let levels = List.map (infos_and_sort env_ar_par) lc in
let isunit = is_unit levels in
@@ -183,7 +184,7 @@ let cumulate_arity_large_levels env sign =
match d with
| LocalAssum (_,t) ->
let tj = infer_type env t in
- let u = univ_of_sort tj.utj_type in
+ let u = Term.univ_of_sort tj.utj_type in
(Universe.sup u lev, push_rel d env)
| LocalDef _ ->
lev, push_rel d env)
@@ -199,8 +200,8 @@ let is_impredicative env u =
let param_ccls paramsctxt =
let fold acc = function
| (LocalAssum (_, p)) ->
- (let c = strip_prod_assum p in
- match kind_of_term c with
+ (let c = Term.strip_prod_assum p in
+ match kind c with
| Sort (Type u) -> Univ.Universe.level u
| _ -> None) :: acc
| LocalDef _ -> acc
@@ -208,7 +209,7 @@ let param_ccls paramsctxt =
List.fold_left fold [] paramsctxt
(* Check arities and constructors *)
-let check_subtyping_arity_constructor env (subst : constr -> constr) (arcn : Term.types) numparams is_arity =
+let check_subtyping_arity_constructor env (subst : constr -> constr) (arcn : types) numparams is_arity =
let numchecked = ref 0 in
let basic_check ev tp =
if !numchecked < numparams then () else conv_leq ev tp (subst tp);
@@ -288,7 +289,7 @@ let typecheck_inductive env mie =
(** We have an algebraic universe as the conclusion of the arity,
typecheck the dummy Π ctx, Prop and do a special case for the conclusion.
*)
- let proparity = infer_type env_params (mkArity (ctx, prop_sort)) in
+ let proparity = infer_type env_params (mkArity (ctx, Sorts.prop)) in
let (cctx, _) = destArity proparity.utj_val in
(* Any universe is well-formed, we don't need to check [s] here *)
mkArity (cctx, s)
@@ -468,7 +469,7 @@ let check_correct_par (env,n,ntypes,_) paramdecls ind_index args =
| LocalDef _ :: paramdecls ->
check param_index (paramdecl_index+1) paramdecls
| _::paramdecls ->
- match kind_of_term (whd_all env params.(param_index)) with
+ match kind (whd_all env params.(param_index)) with
| Rel w when Int.equal w paramdecl_index ->
check (param_index-1) (paramdecl_index+1) paramdecls
| _ ->
@@ -495,7 +496,7 @@ if Int.equal nmr 0 then 0 else
| (_,[]) -> assert false (* |paramsctxt|>=nmr *)
| (lp, LocalDef _ :: paramsctxt) -> find k (index-1) (lp,paramsctxt)
| (p::lp,_::paramsctxt) ->
- ( match kind_of_term (whd_all env p) with
+ ( match kind (whd_all env p) with
| Rel w when Int.equal w index -> find (k+1) (index-1) (lp,paramsctxt)
| _ -> k)
in find 0 (n-1) (lpar,List.rev paramsctxt)
@@ -526,7 +527,7 @@ let ienv_push_inductive (env, n, ntypes, ra_env) ((mi,u),lrecparams) =
let rec ienv_decompose_prod (env,_,_,_ as ienv) n c =
if Int.equal n 0 then (ienv,c) else
let c' = whd_all env c in
- match kind_of_term c' with
+ match kind c' with
Prod(na,a,b) ->
let ienv' = ienv_push_var ienv (na,a,mk_norec) in
ienv_decompose_prod ienv' (n-1) b
@@ -554,8 +555,8 @@ let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (
constructor [cn] has a type of the shape [… -> c … -> P], where,
more generally, the arrows may be dependent). *)
let rec check_pos (env, n, ntypes, ra_env as ienv) nmr c =
- let x,largs = decompose_app (whd_all env c) in
- match kind_of_term x with
+ let x,largs = Term.decompose_app (whd_all env c) in
+ match kind x with
| Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
(** If one of the inductives of the mutually inductive
@@ -662,8 +663,8 @@ let check_positivity_one ~chkpos recursive (env,_,ntypes,_ as ienv) paramsctxt (
inductive type. *)
and check_constructors ienv check_head nmr c =
let rec check_constr_rec (env,n,ntypes,ra_env as ienv) nmr lrec c =
- let x,largs = decompose_app (whd_all env c) in
- match kind_of_term x with
+ let x,largs = Term.decompose_app (whd_all env c) in
+ match kind x with
| Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
@@ -746,7 +747,7 @@ let allowed_sorts is_smashed s =
as well. *)
all_sorts
else
- match family_of_sort s with
+ match Sorts.family s with
(* Type: all elimination allowed: above and below *)
| InType -> all_sorts
(* Smashed Set is necessarily impredicative: forbids large elimination *)
diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli
index f5d0ed3f2..9a9380adb 100644
--- a/kernel/indtypes.mli
+++ b/kernel/indtypes.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Declarations
open Environ
open Entries
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index a39307368..cb03a4d6b 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -10,7 +10,7 @@ open CErrors
open Util
open Names
open Univ
-open Term
+open Constr
open Vars
open Declarations
open Declareops
@@ -29,21 +29,21 @@ let lookup_mind_specif env (kn,tyi) =
(mib, mib.mind_packets.(tyi))
let find_rectype env c =
- let (t, l) = decompose_app (whd_all env c) in
- match kind_of_term t with
+ let (t, l) = Term.decompose_app (whd_all env c) in
+ match kind t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
let find_inductive env c =
- let (t, l) = decompose_app (whd_all env c) in
- match kind_of_term t with
+ let (t, l) = Term.decompose_app (whd_all env c) in
+ match kind t with
| Ind ind
when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite <> Decl_kinds.CoFinite -> (ind, l)
| _ -> raise Not_found
let find_coinductive env c =
- let (t, l) = decompose_app (whd_all env c) in
- match kind_of_term t with
+ let (t, l) = Term.decompose_app (whd_all env c) in
+ match kind t with
| Ind ind
when (fst (lookup_mind_specif env (out_punivs ind))).mind_finite == Decl_kinds.CoFinite -> (ind, l)
| _ -> raise Not_found
@@ -81,7 +81,7 @@ let instantiate_params full t u args sign =
let (rem_args, subs, ty) =
Context.Rel.fold_outside
(fun decl (largs,subs,ty) ->
- match (decl, largs, kind_of_term ty) with
+ match (decl, largs, kind ty) with
| (LocalAssum _, a::args, Prod(_,_,t)) -> (args, a::subs, t)
| (LocalDef (_,b,_), _, LetIn(_,_,_,t)) ->
(largs, (substl subs (subst_instance_constr u b))::subs, t)
@@ -94,9 +94,9 @@ let instantiate_params full t u args sign =
substl subs ty
let full_inductive_instantiate mib u params sign =
- let dummy = prop_sort in
- let t = mkArity (Vars.subst_instance_context u sign,dummy) in
- fst (destArity (instantiate_params true t u params mib.mind_params_ctxt))
+ let dummy = Sorts.prop in
+ let t = Term.mkArity (Vars.subst_instance_context u sign,dummy) in
+ fst (Term.destArity (instantiate_params true t u params mib.mind_params_ctxt))
let full_constructor_instantiate ((mind,_),u,(mib,_),params) t =
let inst_ind = constructor_instantiate mind u mib t in
@@ -128,7 +128,7 @@ where
Remark: Set (predicative) is encoded as Type(0)
*)
-let sort_as_univ = function
+let sort_as_univ = let open Sorts in function
| Type u -> u
| Prop Null -> Universe.type0m
| Prop Pos -> Universe.type0
@@ -192,11 +192,11 @@ let instantiate_universes env ctx ar argsorts =
let level = Univ.subst_univs_universe (Univ.make_subst subst) ar.template_level in
let ty =
(* Singleton type not containing types are interpretable in Prop *)
- if is_type0m_univ level then prop_sort
+ if is_type0m_univ level then Sorts.prop
(* Non singleton type not containing types are interpretable in Set *)
- else if is_type0_univ level then set_sort
+ else if is_type0_univ level then Sorts.set
(* This is a Type with constraints *)
- else Type level
+ else Sorts.Type level
in
(ctx, ty)
@@ -211,9 +211,9 @@ let type_of_inductive_gen ?(polyprop=true) env ((mib,mip),u) paramtyps =
(* The Ocaml extraction cannot handle (yet?) "Prop-polymorphism", i.e.
the situation where a non-Prop singleton inductive becomes Prop
when applied to Prop params *)
- if not polyprop && not (is_type0m_univ ar.template_level) && is_prop_sort s
+ if not polyprop && not (is_type0m_univ ar.template_level) && Sorts.is_prop s
then raise (SingletonInductiveBecomesProp mip.mind_typename);
- mkArity (List.rev ctx,s)
+ Term.mkArity (List.rev ctx,s)
let type_of_inductive env pind =
type_of_inductive_gen env pind [||]
@@ -233,7 +233,7 @@ let type_of_inductive_knowing_parameters env ?(polyprop=true) mip args =
(* The max of an array of universes *)
-let cumulate_constructor_univ u = function
+let cumulate_constructor_univ u = let open Sorts in function
| Prop Null -> u
| Prop Pos -> Universe.sup Universe.type0 u
| Type u' -> Universe.sup u u'
@@ -276,8 +276,8 @@ let type_of_constructors (ind,u) (mib,mip) =
let inductive_sort_family mip =
match mip.mind_arity with
- | RegularArity s -> family_of_sort s.mind_sort
- | TemplateArity _ -> InType
+ | RegularArity s -> Sorts.family s.mind_sort
+ | TemplateArity _ -> Sorts.InType
let mind_arity mip =
mip.mind_arity_ctxt, inductive_sort_family mip
@@ -296,19 +296,20 @@ let is_primitive_record (mib,_) =
let build_dependent_inductive ind (_,mip) params =
let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in
- applist
+ Term.applist
(mkIndU ind,
List.map (lift mip.mind_nrealdecls) params
@ Context.Rel.to_extended_list mkRel 0 realargs)
(* This exception is local *)
-exception LocalArity of (sorts_family * sorts_family * arity_error) option
+exception LocalArity of (Sorts.family * Sorts.family * arity_error) option
let check_allowed_sort ksort specif =
+ let open Sorts in
let eq_ksort s = match ksort, s with
| InProp, InProp | InSet, InSet | InType, InType -> true
| _ -> false in
- if not (List.exists eq_ksort (elim_sorts specif)) then
+ if not (CList.exists eq_ksort (elim_sorts specif)) then
let s = inductive_sort_family (snd specif) in
raise (LocalArity (Some(ksort,s,error_elim_explain ksort s)))
@@ -316,7 +317,7 @@ let is_correct_arity env c pj ind specif params =
let arsign,_ = get_instantiated_arity ind specif params in
let rec srec env pt ar =
let pt' = whd_all env pt in
- match kind_of_term pt', ar with
+ match kind pt', ar with
| Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' ->
let () =
try conv env a1 a1'
@@ -325,8 +326,8 @@ let is_correct_arity env c pj ind specif params =
(* The last Prod domain is the type of the scrutinee *)
| Prod (na1,a1,a2), [] -> (* whnf of t was not needed here! *)
let env' = push_rel (LocalAssum (na1,a1)) env in
- let ksort = match kind_of_term (whd_all env' a2) with
- | Sort s -> family_of_sort s
+ let ksort = match kind (whd_all env' a2) with
+ | Sort s -> Sorts.family s
| _ -> raise (LocalArity None) in
let dep_ind = build_dependent_inductive ind specif params in
let _ =
@@ -351,22 +352,22 @@ let is_correct_arity env c pj ind specif params =
let build_branches_type (ind,u) (_,mip as specif) params p =
let build_one_branch i cty =
let typi = full_constructor_instantiate (ind,u,specif,params) cty in
- let (cstrsign,ccl) = decompose_prod_assum typi in
+ let (cstrsign,ccl) = Term.decompose_prod_assum typi in
let nargs = Context.Rel.length cstrsign in
- let (_,allargs) = decompose_app ccl in
+ let (_,allargs) = Term.decompose_app ccl in
let (lparams,vargs) = List.chop (inductive_params specif) allargs in
let cargs =
let cstr = ith_constructor_of_inductive ind (i+1) in
- let dep_cstr = applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list mkRel 0 cstrsign)) in
+ let dep_cstr = Term.applist (mkConstructU (cstr,u),lparams@(Context.Rel.to_extended_list mkRel 0 cstrsign)) in
vargs @ [dep_cstr] in
- let base = lambda_appvect_assum (mip.mind_nrealdecls+1) (lift nargs p) (Array.of_list cargs) in
- it_mkProd_or_LetIn base cstrsign in
+ let base = Term.lambda_appvect_assum (mip.mind_nrealdecls+1) (lift nargs p) (Array.of_list cargs) in
+ Term.it_mkProd_or_LetIn base cstrsign in
Array.mapi build_one_branch mip.mind_nf_lc
(* [p] is the predicate, [c] is the match object, [realargs] is the
list of real args of the inductive type *)
let build_case_type env n p c realargs =
- whd_betaiota env (lambda_appvect_assum (n+1) p (Array.of_list (realargs@[c])))
+ whd_betaiota env (Term.lambda_appvect_assum (n+1) p (Array.of_list (realargs@[c])))
let type_case_branches env (pind,largs) pj c =
let specif = lookup_mind_specif env (fst pind) in
@@ -565,8 +566,8 @@ let check_inductive_codomain env p =
let env = push_rel_context absctx env in
let arctx, s = dest_prod_assum env ar in
let env = push_rel_context arctx env in
- let i,l' = decompose_app (whd_all env s) in
- isInd i
+ let i,l' = Term.decompose_app (whd_all env s) in
+ Term.isInd i
(* The following functions are almost duplicated from indtypes.ml, except
that they carry here a poorer environment (containing less information). *)
@@ -589,7 +590,7 @@ let ienv_push_inductive (env, ra_env) ((mind,u),lpar) =
let rec ienv_decompose_prod (env,_ as ienv) n c =
if Int.equal n 0 then (ienv,c) else
let c' = whd_all env c in
- match kind_of_term c' with
+ match kind c' with
Prod(na,a,b) ->
let ienv' = ienv_push_var ienv (na,a,mk_norec) in
ienv_decompose_prod ienv' (n-1) b
@@ -620,8 +621,8 @@ close to check_positive in indtypes.ml, but does no positivity check and does no
compute the number of recursive arguments. *)
let get_recargs_approx env tree ind args =
let rec build_recargs (env, ra_env as ienv) tree c =
- let x,largs = decompose_app (whd_all env c) in
- match kind_of_term x with
+ let x,largs = Term.decompose_app (whd_all env c) in
+ match kind x with
| Prod (na,b,d) ->
assert (List.is_empty largs);
build_recargs (ienv_push_var ienv (na, b, mk_norec)) tree d
@@ -679,8 +680,8 @@ let get_recargs_approx env tree ind args =
and build_recargs_constructors ienv trees c =
let rec recargs_constr_rec (env,ra_env as ienv) trees lrec c =
- let x,largs = decompose_app (whd_all env c) in
- match kind_of_term x with
+ let x,largs = Term.decompose_app (whd_all env c) in
+ match kind x with
| Prod (na,b,d) ->
let () = assert (List.is_empty largs) in
@@ -708,8 +709,8 @@ let restrict_spec env spec p =
let env = push_rel_context absctx env in
let arctx, s = dest_prod_assum env ar in
let env = push_rel_context arctx env in
- let i,args = decompose_app (whd_all env s) in
- match kind_of_term i with
+ let i,args = Term.decompose_app (whd_all env s) in
+ match kind i with
| Ind i ->
begin match spec with
| Dead_code -> spec
@@ -729,8 +730,8 @@ let restrict_spec env spec p =
let rec subterm_specif renv stack t =
(* maybe reduction is not always necessary! *)
- let f,l = decompose_app (whd_all renv.env t) in
- match kind_of_term f with
+ let f,l = Term.decompose_app (whd_all renv.env t) in
+ match kind f with
| Rel k -> subterm_var k renv
| Case (ci,p,c,lbr) ->
let stack' = push_stack_closures renv l stack in
@@ -773,7 +774,7 @@ let rec subterm_specif renv stack t =
let decrArg = recindxs.(i) in
let theBody = bodies.(i) in
let nbOfAbst = decrArg+1 in
- let sign,strippedBody = decompose_lam_n_assum nbOfAbst theBody in
+ let sign,strippedBody = Term.decompose_lam_n_assum nbOfAbst theBody in
(* pushing the fix parameters *)
let stack' = push_stack_closures renv l stack in
let renv'' = push_ctxt_renv renv' sign in
@@ -857,13 +858,13 @@ let filter_stack_domain env ci p stack =
else let env = push_rel_context absctx env in
let rec filter_stack env ar stack =
let t = whd_all env ar in
- match stack, kind_of_term t with
+ match stack, kind t with
| elt :: stack', Prod (n,a,c0) ->
let d = LocalAssum (n,a) in
let ctx, a = dest_prod_assum env a in
let env = push_rel_context ctx env in
- let ty, args = decompose_app (whd_all env a) in
- let elt = match kind_of_term ty with
+ let ty, args = Term.decompose_app (whd_all env a) in
+ let elt = match kind ty with
| Ind ind ->
let spec' = stack_element_specif elt in
(match (Lazy.force spec') with
@@ -893,8 +894,8 @@ let check_one_fix renv recpos trees def =
(* if [t] does not make recursive calls, it is guarded: *)
if noccur_with_meta renv.rel_min nfi t then ()
else
- let (f,l) = decompose_app (whd_betaiotazeta renv.env t) in
- match kind_of_term f with
+ let (f,l) = Term.decompose_app (whd_betaiotazeta renv.env t) in
+ match kind f with
| Rel p ->
(* Test if [p] is a fixpoint (recursive call) *)
if renv.rel_min <= p && p < renv.rel_min+nfi then
@@ -924,7 +925,7 @@ let check_one_fix renv recpos trees def =
| LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with FixGuardError _ ->
- check_rec_call renv stack (applist(lift p c,l))
+ check_rec_call renv stack (Term.applist(lift p c,l))
end
| Case (ci,p,c_0,lrest) ->
@@ -970,7 +971,7 @@ let check_one_fix renv recpos trees def =
if evaluable_constant kn renv.env then
try List.iter (check_rec_call renv []) l
with (FixGuardError _ ) ->
- let value = (applist(constant_value_in renv.env cu, l)) in
+ let value = (Term.applist(constant_value_in renv.env cu, l)) in
check_rec_call renv stack value
else List.iter (check_rec_call renv []) l
@@ -1007,7 +1008,7 @@ let check_one_fix renv recpos trees def =
| LocalDef (_,c,_) ->
try List.iter (check_rec_call renv []) l
with (FixGuardError _) ->
- check_rec_call renv stack (applist(c,l))
+ check_rec_call renv stack (Term.applist(c,l))
end
| Sort _ ->
@@ -1022,7 +1023,7 @@ let check_one_fix renv recpos trees def =
if Int.equal decr 0 then
check_rec_call (assign_var_spec renv (1,recArgsDecrArg)) [] body
else
- match kind_of_term body with
+ match kind body with
| Lambda (x,a,b) ->
check_rec_call renv [] a;
let renv' = push_var_renv renv (x,a) in
@@ -1053,7 +1054,7 @@ let inductive_of_mutfix env ((nvect,bodynum),(names,types,bodies as recdef)) =
(* check fi does not appear in the k+1 first abstractions,
gives the type of the k+1-eme abstraction (must be an inductive) *)
let rec check_occur env n def =
- match kind_of_term (whd_all env def) with
+ match kind (whd_all env def) with
| Lambda (x,a,b) ->
if noccur_with_meta n nbfix a then
let env' = push_rel (LocalAssum (x,a)) env in
@@ -1108,7 +1109,7 @@ let anomaly_ill_typed () =
let rec codomain_is_coind env c =
let b = whd_all env c in
- match kind_of_term b with
+ match kind b with
| Prod (x,a,b) ->
codomain_is_coind (push_rel (LocalAssum (x,a)) env) b
| _ ->
@@ -1119,8 +1120,8 @@ let rec codomain_is_coind env c =
let check_one_cofix env nbfix def deftype =
let rec check_rec_call env alreadygrd n tree vlra t =
if not (noccur_with_meta n nbfix t) then
- let c,args = decompose_app (whd_all env t) in
- match kind_of_term c with
+ let c,args = Term.decompose_app (whd_all env t) in
+ match kind c with
| Rel p when n <= p && p < n+nbfix ->
(* recursive call: must be guarded and no nested recursive
call allowed *)
diff --git a/kernel/inductive.mli b/kernel/inductive.mli
index 1caede857..601422a10 100644
--- a/kernel/inductive.mli
+++ b/kernel/inductive.mli
@@ -7,8 +7,8 @@
(************************************************************************)
open Names
-open Term
open Univ
+open Constr
open Declarations
open Environ
@@ -32,23 +32,23 @@ type mind_specif = mutual_inductive_body * one_inductive_body
val lookup_mind_specif : env -> inductive -> mind_specif
(** {6 Functions to build standard types related to inductive } *)
-val ind_subst : MutInd.t -> mutual_inductive_body -> universe_instance -> constr list
+val ind_subst : MutInd.t -> mutual_inductive_body -> Instance.t -> constr list
val inductive_paramdecls : mutual_inductive_body puniverses -> Context.Rel.t
-val instantiate_inductive_constraints :
- mutual_inductive_body -> universe_instance -> constraints
+val instantiate_inductive_constraints :
+ mutual_inductive_body -> Instance.t -> constraints
val constrained_type_of_inductive : env -> mind_specif puniverses -> types constrained
-val constrained_type_of_inductive_knowing_parameters :
+val constrained_type_of_inductive_knowing_parameters :
env -> mind_specif puniverses -> types Lazy.t array -> types constrained
val type_of_inductive : env -> mind_specif puniverses -> types
-val type_of_inductive_knowing_parameters :
+val type_of_inductive_knowing_parameters :
env -> ?polyprop:bool -> mind_specif puniverses -> types Lazy.t array -> types
-val elim_sorts : mind_specif -> sorts_family list
+val elim_sorts : mind_specif -> Sorts.family list
val is_private : mind_specif -> bool
val is_primitive_record : mind_specif -> bool
@@ -85,9 +85,9 @@ val build_branches_type :
constr list -> constr -> types array
(** Return the arity of an inductive type *)
-val mind_arity : one_inductive_body -> Context.Rel.t * sorts_family
+val mind_arity : one_inductive_body -> Context.Rel.t * Sorts.family
-val inductive_sort_family : one_inductive_body -> sorts_family
+val inductive_sort_family : one_inductive_body -> Sorts.family
(** Check a [case_info] actually correspond to a Case expression on the
given inductive type. *)
@@ -111,10 +111,10 @@ val check_cofix : env -> cofixpoint -> unit
exception SingletonInductiveBecomesProp of Id.t
-val max_inductive_sort : sorts array -> universe
+val max_inductive_sort : Sorts.t array -> Universe.t
val instantiate_universes : env -> Context.Rel.t ->
- template_arity -> constr Lazy.t array -> Context.Rel.t * sorts
+ template_arity -> constr Lazy.t array -> Context.Rel.t * Sorts.t
(** {6 Debug} *)
@@ -135,6 +135,6 @@ type stack_element = |SClosure of guard_env*constr |SArg of subterm_spec Lazy.t
val subterm_specif : guard_env -> stack_element list -> constr -> subterm_spec
-val lambda_implicit_lift : int -> Constr.constr -> Term.constr
+val lambda_implicit_lift : int -> constr -> constr
-val abstract_mind_lc : int -> Int.t -> Constr.constr array -> Constr.constr array
+val abstract_mind_lc : int -> Int.t -> constr array -> constr array
diff --git a/kernel/mod_subst.ml b/kernel/mod_subst.ml
index 1793e4f71..2c8ef477f 100644
--- a/kernel/mod_subst.ml
+++ b/kernel/mod_subst.ml
@@ -16,7 +16,7 @@
open Pp
open Util
open Names
-open Term
+open Constr
(* For Inline, the int is an inlining level, and the constr (if present)
is the term into which we should inline. *)
@@ -340,7 +340,7 @@ let subst_evaluable_reference subst = function
let rec map_kn f f' c =
let func = map_kn f f' in
- match kind_of_term c with
+ match kind c with
| Const kn -> (try snd (f' kn) with No_subst -> c)
| Proj (p,t) ->
let p' =
diff --git a/kernel/mod_subst.mli b/kernel/mod_subst.mli
index 1d7012ce5..1aa7ba519 100644
--- a/kernel/mod_subst.mli
+++ b/kernel/mod_subst.mli
@@ -9,7 +9,7 @@
(** {6 [Mod_subst] } *)
open Names
-open Term
+open Constr
(** {6 Delta resolver} *)
@@ -171,6 +171,5 @@ val occur_mbid : MBId.t -> substitution -> bool
val repr_substituted : 'a substituted -> substitution list option * 'a
-val force_constr : Term.constr substituted -> Term.constr
-val subst_constr :
- substitution -> Term.constr substituted -> Term.constr substituted
+val force_constr : constr substituted -> constr
+val subst_constr : substitution -> constr substituted -> constr substituted
diff --git a/kernel/modops.ml b/kernel/modops.ml
index c10af0725..b1df1a187 100644
--- a/kernel/modops.ml
+++ b/kernel/modops.ml
@@ -17,7 +17,7 @@
open Util
open Names
-open Term
+open Constr
open Declarations
open Declareops
open Environ
@@ -266,7 +266,7 @@ let subst_structure subst = subst_structure subst do_delta_codom
(* lclrk : retroknowledge_action list, rkaction : retroknowledge action *)
let add_retroknowledge mp =
let perform rkaction env = match rkaction with
- |Retroknowledge.RKRegister (f, e) when (isConst e || isInd e) ->
+ |Retroknowledge.RKRegister (f, e) when (Term.isConst e || Term.isInd e) ->
Environ.register env f e
|_ ->
CErrors.anomaly ~label:"Modops.add_retroknowledge"
diff --git a/kernel/modops.mli b/kernel/modops.mli
index e1bba21d3..bbb4c918c 100644
--- a/kernel/modops.mli
+++ b/kernel/modops.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Declarations
open Entries
diff --git a/kernel/names.mli b/kernel/names.mli
index 531a0cccb..ba0637c8a 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -300,7 +300,6 @@ module KNset : CSig.SetS with type elt = KerName.t
module KNpred : Predicate.S with type elt = KerName.t
module KNmap : Map.ExtS with type key = KerName.t and module Set := KNset
-
(** {6 Constant Names } *)
module Constant:
diff --git a/kernel/nativecode.ml b/kernel/nativecode.ml
index bb4c2585d..c558e9ed0 100644
--- a/kernel/nativecode.ml
+++ b/kernel/nativecode.ml
@@ -8,7 +8,7 @@
open CErrors
open Names
-open Term
+open Constr
open Declarations
open Util
open Nativevalues
@@ -142,7 +142,7 @@ let fresh_gnormtbl l =
type symbol =
| SymbValue of Nativevalues.t
- | SymbSort of sorts
+ | SymbSort of Sorts.t
| SymbName of Name.t
| SymbConst of Constant.t
| SymbMatch of annot_sw
@@ -163,7 +163,7 @@ let eq_symbol sy1 sy2 =
| SymbInd ind1, SymbInd ind2 -> eq_ind ind1 ind2
| SymbMeta m1, SymbMeta m2 -> Int.equal m1 m2
| SymbEvar (evk1,args1), SymbEvar (evk2,args2) ->
- Evar.equal evk1 evk2 && Array.for_all2 eq_constr args1 args2
+ Evar.equal evk1 evk2 && Array.for_all2 Constr.equal args1 args2
| SymbLevel l1, SymbLevel l2 -> Univ.Level.equal l1 l2
| _, _ -> false
@@ -2016,7 +2016,7 @@ let compile_mind_deps env prefix ~interactive
(* This function compiles all necessary dependencies of t, and generates code in
reverse order, as well as linking information updates *)
let rec compile_deps env sigma prefix ~interactive init t =
- match kind_of_term t with
+ match kind t with
| Ind ((mind,_),u) -> compile_mind_deps env prefix ~interactive init mind
| Const c ->
let c,u = get_alias env c in
@@ -2048,8 +2048,8 @@ let rec compile_deps env sigma prefix ~interactive init t =
| Case (ci, p, c, ac) ->
let mind = fst ci.ci_ind in
let init = compile_mind_deps env prefix ~interactive init mind in
- fold_constr (compile_deps env sigma prefix ~interactive) init t
- | _ -> fold_constr (compile_deps env sigma prefix ~interactive) init t
+ Constr.fold (compile_deps env sigma prefix ~interactive) init t
+ | _ -> Constr.fold (compile_deps env sigma prefix ~interactive) init t
let compile_constant_field env prefix con acc cb =
let (gl, _) =
diff --git a/kernel/nativecode.mli b/kernel/nativecode.mli
index 764598097..d08f49095 100644
--- a/kernel/nativecode.mli
+++ b/kernel/nativecode.mli
@@ -5,8 +5,8 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
open Names
+open Constr
open Declarations
open Pre_env
open Nativelambda
@@ -32,7 +32,7 @@ val clear_symbols : unit -> unit
val get_value : symbols -> int -> Nativevalues.t
-val get_sort : symbols -> int -> sorts
+val get_sort : symbols -> int -> Sorts.t
val get_name : symbols -> int -> Name.t
diff --git a/kernel/nativeconv.mli b/kernel/nativeconv.mli
index fbbcce744..769deacae 100644
--- a/kernel/nativeconv.mli
+++ b/kernel/nativeconv.mli
@@ -5,7 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Reduction
open Nativelambda
diff --git a/kernel/nativeinstr.mli b/kernel/nativeinstr.mli
index b2c8662da..928283a4d 100644
--- a/kernel/nativeinstr.mli
+++ b/kernel/nativeinstr.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Term
+open Constr
open Nativevalues
(** This file defines the lambda code for the native compiler. It has been
@@ -43,7 +43,7 @@ and lambda =
(* A partially applied constructor *)
| Luint of uint
| Lval of Nativevalues.t
- | Lsort of sorts
+ | Lsort of Sorts.t
| Lind of prefix * pinductive
| Llazy
| Lforce
diff --git a/kernel/nativelambda.ml b/kernel/nativelambda.ml
index aaa9f9b00..de4dc2107 100644
--- a/kernel/nativelambda.ml
+++ b/kernel/nativelambda.ml
@@ -8,7 +8,7 @@
open Util
open Names
open Esubst
-open Term
+open Constr
open Declarations
open Pre_env
open Nativevalues
@@ -417,9 +417,9 @@ module Renv =
(* What about pattern matching ?*)
let is_lazy prefix t =
- match kind_of_term t with
+ match kind t with
| App (f,args) ->
- begin match kind_of_term f with
+ begin match kind f with
| Construct (c,_) ->
let entry = mkInd (fst c) in
(try
@@ -448,7 +448,7 @@ let empty_evars =
let empty_ids = [||]
let rec lambda_of_constr env sigma c =
- match kind_of_term c with
+ match kind c with
| Meta mv ->
let ty = meta_type sigma mv in
Lmeta (mv, lambda_of_constr env sigma ty)
@@ -480,7 +480,7 @@ let rec lambda_of_constr env sigma c =
Lprod(ld, Llam([|id|], lc))
| Lambda _ ->
- let params, body = decompose_lam c in
+ let params, body = Term.decompose_lam c in
let ids = get_names (List.rev params) in
Renv.push_rels env ids;
let lb = lambda_of_constr env sigma body in
@@ -561,7 +561,7 @@ let rec lambda_of_constr env sigma c =
Lcofix(init, (names, ltypes, lbodies))
and lambda_of_app env sigma f args =
- match kind_of_term f with
+ match kind f with
| Const (kn,u as c) ->
let kn,u = get_alias !global_env c in
let cb = lookup_constant kn !global_env in
@@ -656,7 +656,7 @@ let compile_static_int31 fc args =
if not fc then raise Not_found else
Luint (UintVal
(Uint31.of_int (Array.fold_left
- (fun temp_i -> fun t -> match kind_of_term t with
+ (fun temp_i -> fun t -> match kind t with
| Construct ((_,d),_) -> 2*temp_i+d-1
| _ -> raise NotClosed)
0 args)))
diff --git a/kernel/nativelambda.mli b/kernel/nativelambda.mli
index 392961f5d..933fbc660 100644
--- a/kernel/nativelambda.mli
+++ b/kernel/nativelambda.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Names
-open Term
+open Constr
open Pre_env
open Nativeinstr
diff --git a/kernel/nativevalues.ml b/kernel/nativevalues.ml
index 2f8920088..ae66362ca 100644
--- a/kernel/nativevalues.ml
+++ b/kernel/nativevalues.ml
@@ -5,10 +5,11 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
-open Names
-open CErrors
+
open Util
+open CErrors
+open Names
+open Constr
(** This module defines the representation of values internally used by
the native compiler *)
@@ -51,7 +52,7 @@ type atom =
| Arel of int
| Aconstant of pconstant
| Aind of pinductive
- | Asort of sorts
+ | Asort of Sorts.t
| Avar of Id.t
| Acase of annot_sw * accumulator * t * (t -> t)
| Afix of t array * t array * rec_pos * int
@@ -111,6 +112,7 @@ let mk_ind_accu ind u =
mk_accu (Aind (ind,Univ.Instance.of_array u))
let mk_sort_accu s u =
+ let open Sorts in
match s with
| Prop _ -> mk_accu (Asort s)
| Type s ->
diff --git a/kernel/nativevalues.mli b/kernel/nativevalues.mli
index 1368b4470..18b877745 100644
--- a/kernel/nativevalues.mli
+++ b/kernel/nativevalues.mli
@@ -5,7 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Names
(** This modules defines the representation of values internally used by
@@ -43,7 +43,7 @@ type atom =
| Arel of int
| Aconstant of pconstant
| Aind of pinductive
- | Asort of sorts
+ | Asort of Sorts.t
| Avar of Id.t
| Acase of annot_sw * accumulator * t * (t -> t)
| Afix of t array * t array * rec_pos * int
@@ -61,7 +61,7 @@ val mk_rel_accu : int -> t
val mk_rels_accu : int -> int -> t array
val mk_constant_accu : Constant.t -> Univ.Level.t array -> t
val mk_ind_accu : inductive -> Univ.Level.t array -> t
-val mk_sort_accu : sorts -> Univ.Level.t array -> t
+val mk_sort_accu : Sorts.t -> Univ.Level.t array -> t
val mk_var_accu : Id.t -> t
val mk_sw_accu : annot_sw -> accumulator -> t -> (t -> t)
val mk_prod_accu : Name.t -> t -> t -> t
diff --git a/kernel/opaqueproof.ml b/kernel/opaqueproof.ml
index 400f9feee..45a62d55a 100644
--- a/kernel/opaqueproof.ml
+++ b/kernel/opaqueproof.ml
@@ -8,7 +8,7 @@
open Names
open Univ
-open Term
+open Constr
open Mod_subst
type work_list = (Instance.t * Id.t array) Cmap.t *
@@ -17,7 +17,7 @@ type work_list = (Instance.t * Id.t array) Cmap.t *
type cooking_info = {
modlist : work_list;
abstract : Context.Named.t * Univ.universe_level_subst * Univ.AUContext.t }
-type proofterm = (constr * Univ.universe_context_set) Future.computation
+type proofterm = (constr * Univ.ContextSet.t) Future.computation
type opaque =
| Indirect of substitution list * DirPath.t * int (* subst, lib, index *)
| Direct of cooking_info list * proofterm
@@ -138,7 +138,7 @@ let get_proof { opaque_val = prfs; opaque_dir = odp } = function
module FMap = Future.UUIDMap
-let a_constr = Future.from_val (Term.mkRel 1)
+let a_constr = Future.from_val (mkRel 1)
let a_univ = Future.from_val Univ.ContextSet.empty
let a_discharge : cooking_info list = []
diff --git a/kernel/opaqueproof.mli b/kernel/opaqueproof.mli
index a0418a022..20d76ce23 100644
--- a/kernel/opaqueproof.mli
+++ b/kernel/opaqueproof.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Mod_subst
(** This module implements the handling of opaque proof terms.
@@ -19,7 +19,7 @@ open Mod_subst
When it is [turn_indirect] the data is relocated to an opaque table
and the [opaque] is turned into an index. *)
-type proofterm = (constr * Univ.universe_context_set) Future.computation
+type proofterm = (constr * Univ.ContextSet.t) Future.computation
type opaquetab
type opaque
@@ -36,10 +36,10 @@ val turn_indirect : DirPath.t -> opaque -> opaquetab -> opaque * opaquetab
(** From a [opaque] back to a [constr]. This might use the
indirect opaque accessor configured below. *)
val force_proof : opaquetab -> opaque -> constr
-val force_constraints : opaquetab -> opaque -> Univ.universe_context_set
-val get_proof : opaquetab -> opaque -> Term.constr Future.computation
+val force_constraints : opaquetab -> opaque -> Univ.ContextSet.t
+val get_proof : opaquetab -> opaque -> constr Future.computation
val get_constraints :
- opaquetab -> opaque -> Univ.universe_context_set Future.computation option
+ opaquetab -> opaque -> Univ.ContextSet.t Future.computation option
val subst_opaque : substitution -> opaque -> opaque
val iter_direct_opaque : (constr -> unit) -> opaque -> opaque
@@ -63,7 +63,7 @@ val join_opaque : opaquetab -> opaque -> unit
val dump : opaquetab ->
Constr.t Future.computation array *
- Univ.universe_context_set Future.computation array *
+ Univ.ContextSet.t Future.computation array *
cooking_info list array *
int Future.UUIDMap.t
@@ -75,7 +75,7 @@ val dump : opaquetab ->
*)
val set_indirect_opaque_accessor :
- (DirPath.t -> int -> Term.constr Future.computation) -> unit
+ (DirPath.t -> int -> constr Future.computation) -> unit
val set_indirect_univ_accessor :
- (DirPath.t -> int -> Univ.universe_context_set Future.computation option) -> unit
+ (DirPath.t -> int -> Univ.ContextSet.t Future.computation option) -> unit
diff --git a/kernel/pre_env.ml b/kernel/pre_env.ml
index 94738d618..c5254b453 100644
--- a/kernel/pre_env.ml
+++ b/kernel/pre_env.ml
@@ -15,7 +15,7 @@
open Util
open Names
-open Term
+open Constr
open Declarations
module NamedDecl = Context.Named.Declaration
diff --git a/kernel/pre_env.mli b/kernel/pre_env.mli
index 27f9511e1..054ae1743 100644
--- a/kernel/pre_env.mli
+++ b/kernel/pre_env.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Declarations
(** The type of environments. *)
diff --git a/kernel/reduction.ml b/kernel/reduction.ml
index 79a35d292..b0510dc7c 100644
--- a/kernel/reduction.ml
+++ b/kernel/reduction.ml
@@ -18,7 +18,7 @@
open CErrors
open Util
open Names
-open Term
+open Constr
open Vars
open Environ
open CClosure
@@ -107,11 +107,11 @@ let pure_stack lfts stk =
(****************************************************************************)
let whd_betaiota env t =
- match kind_of_term t with
+ match kind t with
| (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _|
Prod _|Lambda _|Fix _|CoFix _) -> t
| App (c, _) ->
- begin match kind_of_term c with
+ begin match kind c with
| Ind _ | Construct _ | Evar _ | Meta _ | Const _ | LetIn _ -> t
| _ -> whd_val (create_clos_infos betaiota env) (inject t)
end
@@ -121,33 +121,33 @@ let nf_betaiota env t =
norm_val (create_clos_infos betaiota env) (inject t)
let whd_betaiotazeta env x =
- match kind_of_term x with
+ match kind x with
| (Sort _|Var _|Meta _|Evar _|Const _|Ind _|Construct _|
Prod _|Lambda _|Fix _|CoFix _) -> x
| App (c, _) ->
- begin match kind_of_term c with
+ begin match kind c with
| Ind _ | Construct _ | Evar _ | Meta _ | Const _ -> x
| _ -> whd_val (create_clos_infos betaiotazeta env) (inject x)
end
| _ -> whd_val (create_clos_infos betaiotazeta env) (inject x)
let whd_all env t =
- match kind_of_term t with
+ match kind t with
| (Sort _|Meta _|Evar _|Ind _|Construct _|
Prod _|Lambda _|Fix _|CoFix _) -> t
| App (c, _) ->
- begin match kind_of_term c with
+ begin match kind c with
| Ind _ | Construct _ | Evar _ | Meta _ -> t
| _ -> whd_val (create_clos_infos all env) (inject t)
end
| _ -> whd_val (create_clos_infos all env) (inject t)
let whd_allnolet env t =
- match kind_of_term t with
+ match kind t with
| (Sort _|Meta _|Evar _|Ind _|Construct _|
Prod _|Lambda _|Fix _|CoFix _|LetIn _) -> t
| App (c, _) ->
- begin match kind_of_term c with
+ begin match kind c with
| Ind _ | Construct _ | Evar _ | Meta _ | LetIn _ -> t
| _ -> whd_val (create_clos_infos allnolet env) (inject t)
end
@@ -189,7 +189,7 @@ let is_cumul = function CUMUL -> true | CONV -> false
type 'a universe_compare =
{ (* Might raise NotConvertible *)
- compare : env -> conv_pb -> sorts -> sorts -> 'a -> 'a;
+ compare : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a;
compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
conv_inductives : conv_pb -> (Declarations.mutual_inductive_body * int) -> Univ.Instance.t -> int ->
Univ.Instance.t -> int -> 'a -> 'a;
@@ -331,7 +331,7 @@ and eqappr env cv_pb l2r infos (lft1,st1) (lft2,st2) cuniv =
match (fterm_of hd1, fterm_of hd2) with
(* case of leaves *)
| (FAtom a1, FAtom a2) ->
- (match kind_of_term a1, kind_of_term a2 with
+ (match kind a1, kind a2 with
| (Sort s1, Sort s2) ->
if not (is_empty_stack v1 && is_empty_stack v2) then
anomaly (Pp.str "conversion was given ill-typed terms (Sort).");
@@ -615,6 +615,7 @@ let check_leq univs u u' =
if not (UGraph.check_leq univs u u') then raise NotConvertible
let check_sort_cmp_universes env pb s0 s1 univs =
+ let open Sorts in
match (s0,s1) with
| (Prop c1, Prop c2) when is_cumul pb ->
begin match c1, c2 with
@@ -734,6 +735,7 @@ let infer_leq (univs, cstrs as cuniv) u u' =
univs, cstrs'
let infer_cmp_universes env pb s0 s1 univs =
+ let open Sorts in
match (s0,s1) with
| (Prop c1, Prop c2) when is_cumul pb ->
begin match c1, c2 with
@@ -869,7 +871,7 @@ let warn_bytecode_compiler_failed =
(fun () -> strbrk "Bytecode compiler failed, " ++
strbrk "falling back to standard conversion")
-let set_vm_conv (f:conv_pb -> Term.types kernel_conversion_function) = vm_conv := f
+let set_vm_conv (f:conv_pb -> types kernel_conversion_function) = vm_conv := f
let vm_conv cv_pb env t1 t2 =
try
!vm_conv cv_pb env t1 t2
@@ -895,9 +897,9 @@ let conv env t1 t2 =
let beta_applist c l =
let rec app subst c l =
- match kind_of_term c, l with
+ match kind c, l with
| Lambda(_,_,c), arg::l -> app (arg::subst) c l
- | _ -> applist (substl subst c, l) in
+ | _ -> Term.applist (substl subst c, l) in
app [] c l
let beta_appvect c v = beta_applist c (Array.to_list v)
@@ -905,7 +907,7 @@ let beta_appvect c v = beta_applist c (Array.to_list v)
let beta_app c a = beta_applist c [a]
(* Compatibility *)
-let betazeta_appvect = lambda_appvect_assum
+let betazeta_appvect = Term.lambda_appvect_assum
(********************************************************************)
(* Special-Purpose Reduction *)
@@ -918,7 +920,7 @@ let betazeta_appvect = lambda_appvect_assum
* error message. *)
let hnf_prod_app env t n =
- match kind_of_term (whd_all env t) with
+ match kind (whd_all env t) with
| Prod (_,_,b) -> subst1 n b
| _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product.")
@@ -930,7 +932,7 @@ let hnf_prod_applist env t nl =
let dest_prod env =
let rec decrec env m c =
let t = whd_all env c in
- match kind_of_term t with
+ match kind t with
| Prod (n,a,c0) ->
let d = LocalAssum (n,a) in
decrec (push_rel d env) (Context.Rel.add d m) c0
@@ -942,7 +944,7 @@ let dest_prod env =
let dest_prod_assum env =
let rec prodec_rec env l ty =
let rty = whd_allnolet env ty in
- match kind_of_term rty with
+ match kind rty with
| Prod (x,t,c) ->
let d = LocalAssum (x,t) in
prodec_rec (push_rel d env) (Context.Rel.add d l) c
@@ -952,7 +954,7 @@ let dest_prod_assum env =
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
let rty' = whd_all env rty in
- if Term.eq_constr rty' rty then l, rty
+ if Constr.equal rty' rty then l, rty
else prodec_rec env l rty'
in
prodec_rec env Context.Rel.empty
@@ -960,7 +962,7 @@ let dest_prod_assum env =
let dest_lam_assum env =
let rec lamec_rec env l ty =
let rty = whd_allnolet env ty in
- match kind_of_term rty with
+ match kind rty with
| Lambda (x,t,c) ->
let d = LocalAssum (x,t) in
lamec_rec (push_rel d env) (Context.Rel.add d l) c
@@ -976,7 +978,7 @@ exception NotArity
let dest_arity env c =
let l, c = dest_prod_assum env c in
- match kind_of_term c with
+ match kind c with
| Sort s -> l,s
| _ -> raise NotArity
diff --git a/kernel/reduction.mli b/kernel/reduction.mli
index 253c0874f..05a906e28 100644
--- a/kernel/reduction.mli
+++ b/kernel/reduction.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Environ
(***********************************************************************
@@ -37,7 +37,7 @@ type conv_pb = CONV | CUMUL
type 'a universe_compare =
{ (* Might raise NotConvertible *)
- compare : env -> conv_pb -> sorts -> sorts -> 'a -> 'a;
+ compare : env -> conv_pb -> Sorts.t -> Sorts.t -> 'a -> 'a;
compare_instances: flex:bool -> Univ.Instance.t -> Univ.Instance.t -> 'a -> 'a;
conv_inductives : conv_pb -> (Declarations.mutual_inductive_body * int) -> Univ.Instance.t -> int ->
Univ.Instance.t -> int -> 'a -> 'a;
@@ -51,7 +51,7 @@ type ('a,'b) generic_conversion_function = env -> 'b universe_state -> 'a -> 'a
type 'a infer_conversion_function = env -> UGraph.t -> 'a -> 'a -> Univ.constraints
-val sort_cmp_universes : env -> conv_pb -> sorts -> sorts ->
+val sort_cmp_universes : env -> conv_pb -> Sorts.t -> Sorts.t ->
'a * 'a universe_compare -> 'a * 'a universe_compare
(* [flex] should be true for constants, false for inductive types and
@@ -115,7 +115,7 @@ val dest_lam_assum : env -> types -> Context.Rel.t * types
exception NotArity
-val dest_arity : env -> types -> arity (* raises NotArity if not an arity *)
+val dest_arity : env -> types -> Term.arity (* raises NotArity if not an arity *)
val is_arity : env -> types -> bool
val warn_bytecode_compiler_failed : ?loc:Loc.t -> unit -> unit
diff --git a/kernel/retroknowledge.ml b/kernel/retroknowledge.ml
index 5fbd914f3..88cf93acc 100644
--- a/kernel/retroknowledge.ml
+++ b/kernel/retroknowledge.ml
@@ -14,7 +14,7 @@
for evaluation in the bytecode virtual machine *)
open Names
-open Term
+open Constr
(* The retroknowledge defines a bijective correspondance between some
[entry]-s (which are, in fact, merely terms) and [field]-s which
@@ -102,7 +102,7 @@ module Reactive = Map.Make (EntryOrd)
type reactive_info = {(*information required by the compiler of the VM *)
vm_compiling :
(*fastcomputation flag -> continuation -> result *)
- (bool->Cbytecodes.comp_env->constr array ->
+ (bool -> Cbytecodes.comp_env -> constr array ->
int->Cbytecodes.bytecodes->Cbytecodes.bytecodes)
option;
vm_constant_static :
@@ -117,7 +117,7 @@ type reactive_info = {(*information required by the compiler of the VM *)
(* fastcomputation flag -> cont -> result *)
vm_before_match : (bool -> Cbytecodes.bytecodes -> Cbytecodes.bytecodes) option;
(* tag (= compiled int for instance) -> result *)
- vm_decompile_const : (int -> Term.constr) option;
+ vm_decompile_const : (int -> constr) option;
native_compiling :
(bool -> Nativeinstr.prefix -> Nativeinstr.lambda array ->
diff --git a/kernel/retroknowledge.mli b/kernel/retroknowledge.mli
index 18a12a4ef..e4d78ba14 100644
--- a/kernel/retroknowledge.mli
+++ b/kernel/retroknowledge.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
type retroknowledge
@@ -117,7 +117,7 @@ val get_vm_before_match_info : retroknowledge -> entry -> Cbytecodes.bytecodes
recover the elements of that type from their compiled form if it's non
standard (it is used (and can be used) only when the compiled form
is not a block *)
-val get_vm_decompile_constant_info : retroknowledge -> entry -> int -> Term.constr
+val get_vm_decompile_constant_info : retroknowledge -> entry -> int -> constr
val get_native_compiling_info : retroknowledge -> entry -> Nativeinstr.prefix ->
@@ -163,7 +163,7 @@ type reactive_info = {(*information required by the compiler of the VM *)
(* fastcomputation flag -> cont -> result *)
vm_before_match : (bool -> Cbytecodes.bytecodes -> Cbytecodes.bytecodes) option;
(* tag (= compiled int for instance) -> result *)
- vm_decompile_const : (int -> Term.constr) option;
+ vm_decompile_const : (int -> constr) option;
native_compiling :
(bool -> Nativeinstr.prefix -> Nativeinstr.lambda array ->
diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli
index ee9632944..0bfe07486 100644
--- a/kernel/safe_typing.mli
+++ b/kernel/safe_typing.mli
@@ -69,7 +69,7 @@ val inline_private_constants_in_constr :
val inline_private_constants_in_definition_entry :
Environ.env -> private_constants Entries.definition_entry -> unit Entries.definition_entry
-val universes_of_private : private_constants -> Univ.universe_context_set list
+val universes_of_private : private_constants -> Univ.ContextSet.t list
val is_curmod_library : safe_environment -> bool
@@ -84,13 +84,13 @@ val is_joined_environment : safe_environment -> bool
(** Insertion of local declarations (Local or Variables) *)
val push_named_assum :
- (Id.t * Term.types * bool (* polymorphic *))
+ (Id.t * Constr.types * bool (* polymorphic *))
Univ.in_universe_context_set -> safe_transformer0
(** Returns the full universe context necessary to typecheck the definition
(futures are forced) *)
val push_named_def :
- Id.t * private_constants Entries.definition_entry -> Univ.universe_context_set safe_transformer
+ Id.t * private_constants Entries.definition_entry -> Univ.ContextSet.t safe_transformer
(** Insertion of global axioms or definitions *)
@@ -133,10 +133,10 @@ val add_modtype :
(** Adding universe constraints *)
val push_context_set :
- bool -> Univ.universe_context_set -> safe_transformer0
+ bool -> Univ.ContextSet.t -> safe_transformer0
val push_context :
- bool -> Univ.universe_context -> safe_transformer0
+ bool -> Univ.UContext.t -> safe_transformer0
val add_constraints :
Univ.constraints -> safe_transformer0
@@ -194,18 +194,18 @@ val export :
ModPath.t * compiled_library * native_library
(* Constraints are non empty iff the file is a vi2vo *)
-val import : compiled_library -> Univ.universe_context_set -> vodigest ->
+val import : compiled_library -> Univ.ContextSet.t -> vodigest ->
ModPath.t safe_transformer
(** {6 Safe typing judgments } *)
type judgment
-val j_val : judgment -> Term.constr
-val j_type : judgment -> Term.constr
+val j_val : judgment -> Constr.constr
+val j_type : judgment -> Constr.constr
(** The safe typing of a term returns a typing judgment. *)
-val typing : safe_environment -> Term.constr -> judgment
+val typing : safe_environment -> Constr.constr -> judgment
(** {6 Queries } *)
@@ -221,7 +221,7 @@ open Retroknowledge
val retroknowledge : (retroknowledge-> 'a) -> safe_environment -> 'a
val register :
- field -> Retroknowledge.entry -> Term.constr -> safe_transformer0
+ field -> Retroknowledge.entry -> Constr.constr -> safe_transformer0
val register_inline : Constant.t -> safe_transformer0
diff --git a/kernel/sorts.ml b/kernel/sorts.ml
index cf5207e8d..07688840d 100644
--- a/kernel/sorts.ml
+++ b/kernel/sorts.ml
@@ -14,7 +14,7 @@ type family = InProp | InSet | InType
type t =
| Prop of contents (* proposition types *)
- | Type of universe
+ | Type of Universe.t
let prop = Prop Null
let set = Prop Pos
@@ -91,7 +91,7 @@ module Hsorts =
struct
type _t = t
type t = _t
- type u = universe -> universe
+ type u = Universe.t -> Universe.t
let hashcons huniv = function
| Type u as c ->
diff --git a/kernel/sorts.mli b/kernel/sorts.mli
index 3426d6fd3..65ea75138 100644
--- a/kernel/sorts.mli
+++ b/kernel/sorts.mli
@@ -14,7 +14,7 @@ type family = InProp | InSet | InType
type t =
| Prop of contents (** Prop and Set *)
-| Type of Univ.universe (** Type *)
+| Type of Univ.Universe.t (** Type *)
val set : t
val prop : t
@@ -38,5 +38,5 @@ module List : sig
val intersect : family list -> family list -> family list
end
-val univ_of_sort : t -> Univ.universe
-val sort_of_univ : Univ.universe -> t
+val univ_of_sort : t -> Univ.Universe.t
+val sort_of_univ : Univ.Universe.t -> t
diff --git a/kernel/subtyping.ml b/kernel/subtyping.ml
index d8e281d52..2913c6dfa 100644
--- a/kernel/subtyping.ml
+++ b/kernel/subtyping.ml
@@ -12,10 +12,11 @@
(* This module checks subtyping of module types *)
(*i*)
-open Util
open Names
open Univ
+open Util
open Term
+open Constr
open Declarations
open Declareops
open Reduction
@@ -77,7 +78,7 @@ let make_labmap mp list =
| SFBmodule mb -> { map with mods = Label.Map.add l (Module mb) map.mods }
| SFBmodtype mtb -> { map with mods = Label.Map.add l (Modtype mtb) map.mods }
in
- List.fold_right add_one list empty_labmap
+ CList.fold_right add_one list empty_labmap
let check_conv_error error why cst poly f env a1 a2 =
@@ -153,7 +154,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
let (ctx2,s2) = dest_arity env t2 in
let s1,s2 =
match s1, s2 with
- | Type _, Type _ -> (* shortcut here *) prop_sort, prop_sort
+ | Type _, Type _ -> (* shortcut here *) Sorts.prop, Sorts.prop
| (Prop _, Type _) | (Type _,Prop _) ->
error (NotConvertibleInductiveField name)
| _ -> (s1, s2) in
@@ -216,7 +217,7 @@ let check_inductive cst env mp1 l info1 mp2 mib2 spec2 subst1 subst2 reso1 reso2
(* we check that records and their field names are preserved. *)
check (fun mib -> mib.mind_record <> None) (==) (fun x -> RecordFieldExpected x);
if mib1.mind_record <> None then begin
- let rec names_prod_letin t = match kind_of_term t with
+ let rec names_prod_letin t = match kind t with
| Prod(n,_,t) -> n::(names_prod_letin t)
| LetIn(n,_,_,t) -> n::(names_prod_letin t)
| Cast(t,_,_) -> names_prod_letin t
@@ -272,13 +273,13 @@ let check_constant cst env mp1 l info1 cb2 spec2 subst1 subst2 =
| Type u when not (is_univ_variable u) ->
(* Both types are inferred, no need to recheck them. We
cheat and collapse the types to Prop *)
- mkArity (ctx1,prop_sort), mkArity (ctx2,prop_sort)
+ mkArity (ctx1,Sorts.prop), mkArity (ctx2,Sorts.prop)
| Prop _ ->
(* The type in the interface is inferred, it may be the case
that the type in the implementation is smaller because
the body is more reduced. We safely collapse the upper
type to Prop *)
- mkArity (ctx1,prop_sort), mkArity (ctx2,prop_sort)
+ mkArity (ctx1,Sorts.prop), mkArity (ctx2,Sorts.prop)
| Type _ ->
(* The type in the interface is inferred and the type in the
implementation is not inferred or is inferred but from a
diff --git a/kernel/term.ml b/kernel/term.ml
index 161abfba1..1c970867a 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -20,7 +20,7 @@ type contents = Sorts.contents = Pos | Null
type sorts = Sorts.t =
| Prop of contents (** Prop and Set *)
- | Type of Univ.universe (** Type *)
+ | Type of Univ.Universe.t (** Type *)
type sorts_family = Sorts.family = InProp | InSet | InType
diff --git a/kernel/term.mli b/kernel/term.mli
index 051979180..cb782afac 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Names
+open Constr
(** {5 Redeclaration of types from module Constr and Sorts}
@@ -15,83 +16,7 @@ open Names
*)
-type contents = Sorts.contents = Pos | Null
-
-type sorts = Sorts.t =
- | Prop of contents (** Prop and Set *)
- | Type of Univ.universe (** Type *)
-
-type sorts_family = Sorts.family = InProp | InSet | InType
-
-type 'a puniverses = 'a Univ.puniverses
-
-(** Simply type aliases *)
-type pconstant = Constant.t puniverses
-type pinductive = inductive puniverses
-type pconstructor = constructor puniverses
-
-type constr = Constr.constr
-(** Alias types, for compatibility. *)
-
-type types = Constr.types
-(** Same as [constr], for documentation purposes. *)
-
-type existential_key = Constr.existential_key
-
-type existential = Constr.existential
-
-type metavariable = Constr.metavariable
-
-type case_style = Constr.case_style =
- LetStyle | IfStyle | LetPatternStyle | MatchStyle | RegularStyle
-
-type case_printing = Constr.case_printing =
- { ind_tags : bool list; cstr_tags : bool list array; style : case_style }
-
-type case_info = Constr.case_info =
- { ci_ind : inductive;
- ci_npar : int;
- ci_cstr_ndecls : int array;
- ci_cstr_nargs : int array;
- ci_pp_info : case_printing
- }
-
-type cast_kind = Constr.cast_kind =
- VMcast | NATIVEcast | DEFAULTcast | REVERTcast
-
-type rec_declaration = Constr.rec_declaration
-type fixpoint = Constr.fixpoint
-type cofixpoint = Constr.cofixpoint
-type 'constr pexistential = 'constr Constr.pexistential
-type ('constr, 'types) prec_declaration =
- ('constr, 'types) Constr.prec_declaration
-type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint
-type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint
-
-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 '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.t * '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
-
-type values = Constr.values
-
(** {5 Simple term case analysis. } *)
-
val isRel : constr -> bool
val isRelN : int -> constr -> bool
val isVar : constr -> bool
@@ -118,7 +43,7 @@ val is_Set : constr -> bool
val isprop : constr -> bool
val is_Type : constr -> bool
val iskind : constr -> bool
-val is_small : sorts -> bool
+val is_small : Sorts.t -> bool
(** {5 Term destructors } *)
@@ -138,7 +63,7 @@ val destVar : constr -> Id.t
(** Destructs a sort. [is_Prop] recognizes the sort {% \textsf{%}Prop{% }%}, whether
[isprop] recognizes both {% \textsf{%}Prop{% }%} and {% \textsf{%}Set{% }%}. *)
-val destSort : constr -> sorts
+val destSort : constr -> Sorts.t
(** Destructs a casted term *)
val destCast : constr -> constr * cast_kind * constr
@@ -197,7 +122,6 @@ val destFix : constr -> fixpoint
val destCoFix : constr -> cofixpoint
-
(** {5 Derived constructors} *)
(** non-dependent product [t1 -> t2], an alias for
@@ -354,7 +278,7 @@ val strip_lam_assum : constr -> constr
Such a term can canonically be seen as the pair of a context of types
and of a sort *)
-type arity = Context.Rel.t * sorts
+type arity = Context.Rel.t * Sorts.t
(** Build an "arity" from its canonical form *)
val mkArity : arity -> types
@@ -368,7 +292,7 @@ val isArity : types -> bool
(** {5 Kind of type} *)
type ('constr, 'types) kind_of_type =
- | SortType of sorts
+ | SortType of Sorts.t
| CastType of 'types * 'types
| ProdType of Name.t * 'types * 'types
| LetInType of Name.t * 'constr * 'types * 'types
@@ -378,23 +302,23 @@ val kind_of_type : types -> (constr, types) kind_of_type
(** {5 Redeclaration of stuff from module [Sorts]} *)
-val set_sort : sorts
-(** Alias for Sorts.set *)
+val set_sort : Sorts.t
+[@@ocaml.deprecated "Alias for Sorts.set"]
-val prop_sort : sorts
-(** Alias for Sorts.prop *)
+val prop_sort : Sorts.t
+[@@ocaml.deprecated "Alias for Sorts.prop"]
-val type1_sort : sorts
-(** Alias for Sorts.type1 *)
+val type1_sort : Sorts.t
+[@@ocaml.deprecated "Alias for Sorts.type1"]
-val sorts_ord : sorts -> sorts -> int
-(** Alias for Sorts.compare *)
+val sorts_ord : Sorts.t -> Sorts.t -> int
+[@@ocaml.deprecated "Alias for Sorts.compare"]
-val is_prop_sort : sorts -> bool
-(** Alias for Sorts.is_prop *)
+val is_prop_sort : Sorts.t -> bool
+[@@ocaml.deprecated "Alias for Sorts.is_prop"]
-val family_of_sort : sorts -> sorts_family
-(** Alias for Sorts.family *)
+val family_of_sort : Sorts.t -> Sorts.family
+[@@ocaml.deprecated "Alias for Sorts.family"]
(** {5 Redeclaration of stuff from module [Constr]}
@@ -403,90 +327,212 @@ val family_of_sort : sorts -> sorts_family
(** {6 Term constructors. } *)
val mkRel : int -> constr
+[@@ocaml.deprecated "Alias for Constr.mkRel"]
val mkVar : Id.t -> constr
+[@@ocaml.deprecated "Alias for Constr.mkVar"]
val mkMeta : metavariable -> constr
+[@@ocaml.deprecated "Alias for Constr.mkMeta"]
val mkEvar : existential -> constr
-val mkSort : sorts -> types
+[@@ocaml.deprecated "Alias for Constr.mkEvar"]
+val mkSort : Sorts.t -> types
+[@@ocaml.deprecated "Alias for Constr.mkSort"]
val mkProp : types
+[@@ocaml.deprecated "Alias for Constr.mkProp"]
val mkSet : types
-val mkType : Univ.universe -> types
+[@@ocaml.deprecated "Alias for Constr.mkSet"]
+val mkType : Univ.Universe.t -> types
+[@@ocaml.deprecated "Alias for Constr.mkType"]
val mkCast : constr * cast_kind * constr -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkProd : Name.t * types * types -> types
+[@@ocaml.deprecated "Alias for Constr"]
val mkLambda : Name.t * types * constr -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkLetIn : Name.t * constr * types * constr -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkApp : constr * constr array -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkConst : Constant.t -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkProj : projection * constr -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkInd : inductive -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkConstruct : constructor -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkConstU : Constant.t puniverses -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkIndU : inductive puniverses -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkConstructU : constructor puniverses -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkConstructUi : (pinductive * int) -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkCase : case_info * constr * constr * constr array -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkFix : fixpoint -> constr
+[@@ocaml.deprecated "Alias for Constr"]
val mkCoFix : cofixpoint -> constr
+[@@ocaml.deprecated "Alias for Constr"]
(** {6 Aliases} *)
val eq_constr : constr -> constr -> bool
-(** Alias for [Constr.equal] *)
+[@@ocaml.deprecated "Alias for Constr.equal"]
(** [eq_constr_univs u a b] is [true] if [a] equals [b] modulo alpha, casts,
application grouping and the universe constraints in [u]. *)
val eq_constr_univs : constr UGraph.check_function
+[@@ocaml.deprecated "Alias for Constr.eq_constr_univs"]
(** [leq_constr_univs u a b] is [true] if [a] is convertible to [b] modulo
alpha, casts, application grouping and the universe constraints in [u]. *)
val leq_constr_univs : constr UGraph.check_function
+[@@ocaml.deprecated "Alias for Constr.leq_constr_univs"]
(** [eq_constr_univs a b] [true, c] if [a] equals [b] modulo alpha, casts,
application grouping and ignoring universe instances. *)
val eq_constr_nounivs : constr -> constr -> bool
+[@@ocaml.deprecated "Alias for Constr.qe_constr_nounivs"]
val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
-(** Alias for [Constr.kind] *)
+[@@ocaml.deprecated "Alias for Constr.kind"]
val compare : constr -> constr -> int
-(** Alias for [Constr.compare] *)
+[@@ocaml.deprecated "Alias for [Constr.compare]"]
val constr_ord : constr -> constr -> int
-(** Alias for [Term.compare] *)
+[@@ocaml.deprecated "Alias for [Term.compare]"]
val fold_constr : ('a -> constr -> 'a) -> 'a -> constr -> 'a
-(** Alias for [Constr.fold] *)
+[@@ocaml.deprecated "Alias for [Constr.fold]"]
val map_constr : (constr -> constr) -> constr -> constr
-(** Alias for [Constr.map] *)
+[@@ocaml.deprecated "Alias for [Constr.map]"]
val map_constr_with_binders :
('a -> 'a) -> ('a -> constr -> constr) -> 'a -> constr -> constr
-(** Alias for [Constr.map_with_binders] *)
+[@@ocaml.deprecated "Alias for [Constr.map_with_binders]"]
val map_puniverses : ('a -> 'b) -> 'a puniverses -> 'b puniverses
-val univ_of_sort : sorts -> Univ.universe
-val sort_of_univ : Univ.universe -> sorts
+val univ_of_sort : Sorts.t -> Univ.Universe.t
+val sort_of_univ : Univ.Universe.t -> Sorts.t
val iter_constr : (constr -> unit) -> constr -> unit
-(** Alias for [Constr.iter] *)
+[@@ocaml.deprecated "Alias for [Constr.iter]"]
val iter_constr_with_binders :
('a -> 'a) -> ('a -> constr -> unit) -> 'a -> constr -> unit
-(** Alias for [Constr.iter_with_binders] *)
+[@@ocaml.deprecated "Alias for [Constr.iter_with_binders]"]
val compare_constr : (constr -> constr -> bool) -> constr -> constr -> bool
-(** Alias for [Constr.compare_head] *)
+[@@ocaml.deprecated "Alias for [Constr.compare_head]"]
+
+type constr = Constr.constr
+[@@ocaml.deprecated "Alias for Constr.t"]
+
+(** Alias types, for compatibility. *)
+
+type types = Constr.types
+[@@ocaml.deprecated "Alias for Constr.types"]
+
+type contents = Sorts.contents = Pos | Null
+[@@ocaml.deprecated "Alias for Sorts.contents"]
+
+type sorts = Sorts.t =
+ | Prop of Sorts.contents (** Prop and Set *)
+ | Type of Univ.Universe.t (** Type *)
+[@@ocaml.deprecated "Alias for Sorts.t"]
+
+type sorts_family = Sorts.family = InProp | InSet | InType
+[@@ocaml.deprecated "Alias for Sorts.family"]
+
+type 'a puniverses = 'a Constr.puniverses
+[@@ocaml.deprecated "Alias for Constr.puniverses"]
+
+(** Simply type aliases *)
+type pconstant = Constr.pconstant
+[@@ocaml.deprecated "Alias for Constr.pconstant"]
+type pinductive = Constr.pinductive
+[@@ocaml.deprecated "Alias for Constr.pinductive"]
+type pconstructor = Constr.pconstructor
+[@@ocaml.deprecated "Alias for Constr.pconstructor"]
+type existential_key = Constr.existential_key
+[@@ocaml.deprecated "Alias for Constr.existential_key"]
+type existential = Constr.existential
+[@@ocaml.deprecated "Alias for Constr.existential"]
+type metavariable = Constr.metavariable
+[@@ocaml.deprecated "Alias for Constr.metavariable"]
+
+type case_style = Constr.case_style =
+ LetStyle | IfStyle | LetPatternStyle | MatchStyle | RegularStyle
+[@@ocaml.deprecated "Alias for Constr.case_style"]
+
+type case_printing = Constr.case_printing =
+ { ind_tags : bool list; cstr_tags : bool list array; style : Constr.case_style }
+[@@ocaml.deprecated "Alias for Constr.case_printing"]
+
+type case_info = Constr.case_info =
+ { ci_ind : inductive;
+ ci_npar : int;
+ ci_cstr_ndecls : int array;
+ ci_cstr_nargs : int array;
+ ci_pp_info : Constr.case_printing
+ }
+[@@ocaml.deprecated "Alias for Constr.case_info"]
+
+type cast_kind = Constr.cast_kind =
+ VMcast | NATIVEcast | DEFAULTcast | REVERTcast
+[@@ocaml.deprecated "Alias for Constr.cast_kind"]
+
+type rec_declaration = Constr.rec_declaration
+[@@ocaml.deprecated "Alias for Constr.rec_declaration"]
+type fixpoint = Constr.fixpoint
+[@@ocaml.deprecated "Alias for Constr.fixpoint"]
+type cofixpoint = Constr.cofixpoint
+[@@ocaml.deprecated "Alias for Constr.cofixpoint"]
+type 'constr pexistential = 'constr Constr.pexistential
+[@@ocaml.deprecated "Alias for Constr.pexistential"]
+type ('constr, 'types) prec_declaration =
+ ('constr, 'types) Constr.prec_declaration
+[@@ocaml.deprecated "Alias for Constr.prec_declaration"]
+type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint
+[@@ocaml.deprecated "Alias for Constr.pfixpoint"]
+type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint
+[@@ocaml.deprecated "Alias for Constr.pcofixpoint"]
+
+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 Constr.metavariable
+ | Evar of 'constr Constr.pexistential
+ | Sort of 'sort
+ | Cast of 'constr * 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.t * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
+ | Case of Constr.case_info * 'constr * 'constr * 'constr array
+ | Fix of ('constr, 'types) Constr.pfixpoint
+ | CoFix of ('constr, 'types) Constr.pcofixpoint
+ | Proj of projection * 'constr
+[@@ocaml.deprecated "Alias for Constr.kind_of_term"]
-val hash_constr : constr -> int
-(** Alias for [Constr.hash] *)
+type values = Constr.values
+[@@ocaml.deprecated "Alias for Constr.values"]
-(*********************************************************************)
+val hash_constr : Constr.constr -> int
+[@@ocaml.deprecated "Alias for Constr.hash"]
-val hcons_sorts : sorts -> sorts
-(** Alias for [Constr.hashcons_sorts] *)
+val hcons_sorts : Sorts.t -> Sorts.t
+[@@ocaml.deprecated "Alias for [Sorts.hcons]"]
-val hcons_constr : constr -> constr
-(** Alias for [Constr.hashcons] *)
+val hcons_constr : Constr.constr -> Constr.constr
+[@@ocaml.deprecated "Alias for [Constr.hcons]"]
-val hcons_types : types -> types
-(** Alias for [Constr.hashcons] *)
+val hcons_types : Constr.types -> Constr.types
+[@@ocaml.deprecated "Alias for [Constr.hcons]"]
diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml
index d50abf071..4617f2d5f 100644
--- a/kernel/term_typing.ml
+++ b/kernel/term_typing.ml
@@ -15,7 +15,7 @@
open CErrors
open Util
open Names
-open Term
+open Constr
open Declarations
open Environ
open Entries
@@ -154,7 +154,7 @@ let inline_side_effects env body ctx side_eff =
(** Lift free rel variables *)
if n <= k then t
else mkRel (n + len - i - 1)
- | _ -> map_constr_with_binders ((+) 1) (fun k t -> subst_const i k t) k t
+ | _ -> Constr.map_with_binders ((+) 1) (fun k t -> subst_const i k t) k t
in
let map_args i (na, b, ty, opaque) =
(** Both the type and the body may mention other constants *)
@@ -199,13 +199,13 @@ let check_signatures curmb sl =
let skip_trusted_seff sl b e =
let rec aux sl b e acc =
let open Context.Rel.Declaration in
- match sl, kind_of_term b with
+ match sl, kind b with
| (None|Some 0), _ -> b, e, acc
| Some sl, LetIn (n,c,ty,bo) ->
aux (Some (sl-1)) bo
(Environ.push_rel (LocalDef (n,c,ty)) e) (`Let(n,c,ty)::acc)
| Some sl, App(hd,arg) ->
- begin match kind_of_term hd with
+ begin match kind hd with
| Lambda (n,ty,bo) ->
aux (Some (sl-1)) bo
(Environ.push_rel (LocalAssum (n,ty)) e) (`Cut(n,ty,arg)::acc)
@@ -245,7 +245,7 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry
abstract_constant_universes abstract uctx
in
let c = Typeops.assumption_of_judgment env j in
- let t = hcons_constr (Vars.subst_univs_level_constr usubst c) in
+ let t = Constr.hcons (Vars.subst_univs_level_constr usubst c) in
{
Cooking.cook_body = Undef nl;
cook_type = t;
@@ -283,7 +283,7 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry
let _ = judge_of_cast env j DEFAULTcast tyj in
j, uctx
in
- let c = hcons_constr j.uj_val in
+ let c = Constr.hcons j.uj_val in
feedback_completion_typecheck feedback_id;
c, uctx) in
let def = OpaqueDef (Opaqueproof.create proofterm) in
@@ -325,7 +325,7 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry
let _ = judge_of_cast env j DEFAULTcast tj in
Vars.subst_univs_level_constr usubst t
in
- let def = hcons_constr (Vars.subst_univs_level_constr usubst j.uj_val) in
+ let def = Constr.hcons (Vars.subst_univs_level_constr usubst j.uj_val) in
let def =
if opaque then OpaqueDef (Opaqueproof.create (Future.from_val (def, Univ.ContextSet.empty)))
else Def (Mod_subst.from_val def)
@@ -359,7 +359,7 @@ let infer_declaration (type a) ~(trust : a trust) env kn (dcl : a constant_entry
in
let term, typ = pb.proj_eta in
{
- Cooking.cook_body = Def (Mod_subst.from_val (hcons_constr term));
+ Cooking.cook_body = Def (Mod_subst.from_val (Constr.hcons term));
cook_type = typ;
cook_proj = Some pb;
cook_universes = univs;
diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli
index 815269169..9b35bfc6e 100644
--- a/kernel/term_typing.mli
+++ b/kernel/term_typing.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Declarations
open Entries
@@ -19,7 +19,7 @@ type _ trust =
| SideEffects : structure_body -> side_effects trust
val translate_local_def : 'a trust -> env -> Id.t -> 'a definition_entry ->
- constant_def * types * Univ.universe_context
+ constant_def * types * Univ.UContext.t
val translate_local_assum : env -> types -> types
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 9813fc566..3a1f2ae00 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Reduction
@@ -45,8 +45,8 @@ type ('constr, 'types) ptype_error =
| NotAType of ('constr, 'types) punsafe_judgment
| BadAssumption of ('constr, 'types) punsafe_judgment
| ReferenceVariables of Id.t * 'constr
- | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment
- * (sorts_family * sorts_family * arity_error) option
+ | ElimArity of pinductive * Sorts.family list * 'constr * ('constr, 'types) punsafe_judgment
+ * (Sorts.family * Sorts.family * arity_error) option
| CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
| NumberBranches of ('constr, 'types) punsafe_judgment * int
@@ -115,6 +115,7 @@ let error_ill_typed_rec_body env i lna vdefj vargs =
raise (TypeError (env, IllTypedRecBody (i,lna,vdefj,vargs)))
let error_elim_explain kp ki =
+ let open Sorts in
match kp,ki with
| (InType | InSet), InProp -> NonInformativeToInformative
| InType, InSet -> StrongEliminationOnNonSmallType (* if Set impredicative *)
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index 95a963da2..e4fa65686 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
(** Type errors. {% \label{%}typeerrors{% }%} *)
@@ -46,8 +46,8 @@ type ('constr, 'types) ptype_error =
| NotAType of ('constr, 'types) punsafe_judgment
| BadAssumption of ('constr, 'types) punsafe_judgment
| ReferenceVariables of Id.t * 'constr
- | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment
- * (sorts_family * sorts_family * arity_error) option
+ | ElimArity of pinductive * Sorts.family list * 'constr * ('constr, 'types) punsafe_judgment
+ * (Sorts.family * Sorts.family * arity_error) option
| CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
| NumberBranches of ('constr, 'types) punsafe_judgment * int
@@ -77,8 +77,8 @@ val error_assumption : env -> unsafe_judgment -> 'a
val error_reference_variables : env -> Id.t -> constr -> 'a
val error_elim_arity :
- env -> pinductive -> sorts_family list -> constr -> unsafe_judgment ->
- (sorts_family * sorts_family * arity_error) option -> 'a
+ env -> pinductive -> Sorts.family list -> constr -> unsafe_judgment ->
+ (Sorts.family * Sorts.family * arity_error) option -> 'a
val error_case_not_inductive : env -> unsafe_judgment -> 'a
@@ -103,6 +103,6 @@ val error_ill_formed_rec_body :
val error_ill_typed_rec_body :
env -> int -> Name.t array -> unsafe_judgment array -> types array -> 'a
-val error_elim_explain : sorts_family -> sorts_family -> arity_error
+val error_elim_explain : Sorts.family -> Sorts.family -> arity_error
val error_unsatisfied_constraints : env -> Univ.constraints -> 'a
diff --git a/kernel/typeops.ml b/kernel/typeops.ml
index f99f34b06..4ccef5c38 100644
--- a/kernel/typeops.ml
+++ b/kernel/typeops.ml
@@ -10,7 +10,8 @@ open CErrors
open Util
open Names
open Univ
-open Term
+open Sorts
+open Constr
open Vars
open Declarations
open Environ
@@ -38,7 +39,7 @@ let check_constraints cst env =
(* This should be a type (a priori without intention to be an assumption) *)
let check_type env c t =
- match kind_of_term(whd_all env t) with
+ match kind(whd_all env t) with
| Sort s -> s
| _ -> error_not_type env (make_judge c t)
@@ -57,7 +58,7 @@ let check_assumption env t ty =
(* Prop and Set *)
-let type1 = mkSort type1_sort
+let type1 = mkSort Sorts.type1
(* Type of Type(i). *)
@@ -152,7 +153,7 @@ let type_of_apply env func funt argsv argstv =
let rec apply_rec i typ =
if Int.equal i len then typ
else
- (match kind_of_term (whd_all env typ) with
+ (match kind (whd_all env typ) with
| Prod (_,c1,c2) ->
let arg = argsv.(i) and argt = argstv.(i) in
(try
@@ -300,7 +301,7 @@ let type_of_projection env p c ct =
in
assert(MutInd.equal pb.proj_ind (fst ind));
let ty = Vars.subst_instance_constr u pb.Declarations.proj_type in
- substl (c :: List.rev args) ty
+ substl (c :: CList.rev args) ty
(* Fixpoints. *)
@@ -325,7 +326,7 @@ let check_fixpoint env lna lar vdef vdeft =
arbitraires et non plus des variables *)
let rec execute env cstr =
let open Context.Rel.Declaration in
- match kind_of_term cstr with
+ match kind cstr with
(* Atomic terms *)
| Sort s -> type_of_sort s
@@ -346,7 +347,7 @@ let rec execute env cstr =
| App (f,args) ->
let argst = execute_array env args in
let ft =
- match kind_of_term f with
+ match kind f with
| Ind ind when Environ.template_polymorphic_pind ind env ->
let args = Array.map (fun t -> lazy t) argst in
type_of_inductive_knowing_parameters env ind args
diff --git a/kernel/typeops.mli b/kernel/typeops.mli
index 96be6c14a..3aaad5877 100644
--- a/kernel/typeops.mli
+++ b/kernel/typeops.mli
@@ -8,7 +8,7 @@
open Names
open Univ
-open Term
+open Constr
open Environ
open Entries
@@ -41,8 +41,8 @@ val type1 : types
val type_of_sort : Sorts.t -> types
val judge_of_prop : unsafe_judgment
val judge_of_set : unsafe_judgment
-val judge_of_prop_contents : contents -> unsafe_judgment
-val judge_of_type : universe -> unsafe_judgment
+val judge_of_prop_contents : Sorts.contents -> unsafe_judgment
+val judge_of_type : Universe.t -> unsafe_judgment
(** {6 Type of a bound variable. } *)
val type_of_relative : env -> int -> types
@@ -71,8 +71,8 @@ val judge_of_abstraction :
-> unsafe_judgment
(** {6 Type of a product. } *)
-val sort_of_product : env -> sorts -> sorts -> sorts
-val type_of_product : env -> Name.t -> sorts -> sorts -> types
+val sort_of_product : env -> Sorts.t -> Sorts.t -> Sorts.t
+val type_of_product : env -> Name.t -> Sorts.t -> Sorts.t -> types
val judge_of_product :
env -> Name.t -> unsafe_type_judgment -> unsafe_type_judgment
-> unsafe_judgment
diff --git a/kernel/uGraph.ml b/kernel/uGraph.ml
index 9793dd881..00c0ea70d 100644
--- a/kernel/uGraph.ml
+++ b/kernel/uGraph.ml
@@ -910,3 +910,4 @@ let check_leq =
let check_leq_key = Profile.declare_profile "check_leq" in
Profile.profile3 check_leq_key check_leq
else check_leq
+
diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli
index 2fe555018..b95388ed0 100644
--- a/kernel/uGraph.mli
+++ b/kernel/uGraph.mli
@@ -9,63 +9,64 @@
open Univ
(** {6 Graphs of universes. } *)
-
type t
-
type universes = t
+[@@ocaml.deprecated "Use UGraph.t"]
-type 'a check_function = universes -> 'a -> 'a -> bool
-val check_leq : universe check_function
-val check_eq : universe check_function
-val check_eq_level : universe_level check_function
+type 'a check_function = t -> 'a -> 'a -> bool
-(** The empty graph of universes *)
-val empty_universes : universes
+val check_leq : Universe.t check_function
+val check_eq : Universe.t check_function
+val check_eq_level : Level.t check_function
(** The initial graph of universes: Prop < Set *)
-val initial_universes : universes
+val initial_universes : t
+
+(** Check if we are in the initial case *)
+val is_initial_universes : t -> bool
+
+(** Check equality of instances w.r.t. a universe graph *)
+val check_eq_instances : Instance.t check_function
+
+(** {6 ... } *)
+(** Merge of constraints in a universes graph.
+ The function [merge_constraints] merges a set of constraints in a given
+ universes graph. It raises the exception [UniverseInconsistency] if the
+ constraints are not satisfiable. *)
-val is_initial_universes : universes -> bool
+val enforce_constraint : univ_constraint -> t -> t
+val merge_constraints : constraints -> t -> t
-val sort_universes : universes -> universes
+val check_constraint : t -> univ_constraint -> bool
+val check_constraints : constraints -> t -> bool
(** Adds a universe to the graph, ensuring it is >= or > Set.
@raises AlreadyDeclared if the level is already declared in the graph. *)
exception AlreadyDeclared
-val add_universe : universe_level -> bool -> universes -> universes
+val add_universe : Level.t -> bool -> t -> t
-(** {6 ... } *)
-(** Merge of constraints in a universes graph.
- The function [merge_constraints] merges a set of constraints in a given
- universes graph. It raises the exception [UniverseInconsistency] if the
- constraints are not satisfiable. *)
+(** {6 Pretty-printing of universes. } *)
-val enforce_constraint : univ_constraint -> universes -> universes
-val merge_constraints : constraints -> universes -> universes
+val pr_universes : (Level.t -> Pp.t) -> t -> Pp.t
-val constraints_of_universes : universes -> constraints
+(** The empty graph of universes *)
+val empty_universes : t
+[@@ocaml.deprecated "Use UGraph.initial_universes"]
-val check_constraint : universes -> univ_constraint -> bool
-val check_constraints : constraints -> universes -> bool
+val sort_universes : t -> t
-val check_eq_instances : Instance.t check_function
-(** Check equality of instances w.r.t. a universe graph *)
+val constraints_of_universes : t -> constraints
val check_subtype : AUContext.t check_function
(** [check_subtype univ ctx1 ctx2] checks whether [ctx2] is an instance of
[ctx1]. *)
-(** {6 Pretty-printing of universes. } *)
-
-val pr_universes : (Level.t -> Pp.t) -> universes -> Pp.t
-
(** {6 Dumping to a file } *)
val dump_universes :
- (constraint_type -> string -> string -> unit) ->
- universes -> unit
+ (constraint_type -> string -> string -> unit) -> t -> unit
(** {6 Debugging} *)
-val check_universes_invariants : universes -> unit
+val check_universes_invariants : t -> unit
diff --git a/kernel/univ.mli b/kernel/univ.mli
index 94116e473..8d46a8bee 100644
--- a/kernel/univ.mli
+++ b/kernel/univ.mli
@@ -7,7 +7,6 @@
(************************************************************************)
(** Universes. *)
-
module Level :
sig
type t
@@ -20,11 +19,11 @@ sig
val is_small : t -> bool
(** Is the universe set or prop? *)
-
+
val is_prop : t -> bool
val is_set : t -> bool
(** Is it specifically Prop or Set *)
-
+
val compare : t -> t -> int
(** Comparison function *)
@@ -49,18 +48,19 @@ sig
end
type universe_level = Level.t
-(** Alias name. *)
+[@@ocaml.deprecated "Use Level.t"]
(** Sets of universe levels *)
-module LSet :
-sig
- include CSig.SetS with type elt = universe_level
-
+module LSet :
+sig
+ include CSig.SetS with type elt = Level.t
+
val pr : (Level.t -> Pp.t) -> t -> Pp.t
(** Pretty-printing *)
end
type universe_set = LSet.t
+[@@ocaml.deprecated "Use LSet.t"]
module Universe :
sig
@@ -106,17 +106,17 @@ sig
val super : t -> t
(** The universe strictly above *)
-
+
val sup : t -> t -> t
(** The l.u.b. of 2 universes *)
- val type0m : t
+ val type0m : t
(** image of Prop in the universes hierarchy *)
-
- val type0 : t
+
+ val type0 : t
(** image of Set in the universes hierarchy *)
-
- val type1 : t
+
+ val type1 : t
(** the universe of the type of Prop/Set *)
val exists : (Level.t * int -> bool) -> t -> bool
@@ -124,40 +124,41 @@ sig
end
type universe = Universe.t
+[@@ocaml.deprecated "Use Universe.t"]
(** Alias name. *)
-val pr_uni : universe -> Pp.t
-
-(** The universes hierarchy: Type 0- = Prop <= Type 0 = Set <= Type 1 <= ...
+val pr_uni : Universe.t -> Pp.t
+
+(** The universes hierarchy: Type 0- = Prop <= Type 0 = Set <= Type 1 <= ...
Typing of universes: Type 0-, Type 0 : Type 1; Type i : Type (i+1) if i>0 *)
-val type0m_univ : universe
-val type0_univ : universe
-val type1_univ : universe
+val type0m_univ : Universe.t
+val type0_univ : Universe.t
+val type1_univ : Universe.t
-val is_type0_univ : universe -> bool
-val is_type0m_univ : universe -> bool
-val is_univ_variable : universe -> bool
-val is_small_univ : universe -> bool
+val is_type0_univ : Universe.t -> bool
+val is_type0m_univ : Universe.t -> bool
+val is_univ_variable : Universe.t -> bool
+val is_small_univ : Universe.t -> bool
-val sup : universe -> universe -> universe
-val super : universe -> universe
+val sup : Universe.t -> Universe.t -> Universe.t
+val super : Universe.t -> Universe.t
-val universe_level : universe -> universe_level option
+val universe_level : Universe.t -> Level.t option
(** [univ_level_mem l u] Is l is mentionned in u ? *)
-val univ_level_mem : universe_level -> universe -> bool
+val univ_level_mem : Level.t -> Universe.t -> bool
(** [univ_level_rem u v min] removes [u] from [v], resulting in [min]
if [v] was exactly [u]. *)
-val univ_level_rem : universe_level -> universe -> universe -> universe
+val univ_level_rem : Level.t -> Universe.t -> Universe.t -> Universe.t
(** {6 Constraints. } *)
type constraint_type = Lt | Le | Eq
-type univ_constraint = universe_level * constraint_type * universe_level
+type univ_constraint = Level.t * constraint_type * Level.t
module Constraint : sig
include Set.S with type elt = univ_constraint
@@ -179,10 +180,10 @@ val constraints_of : 'a constrained -> constraints
type 'a constraint_function = 'a -> 'a -> constraints -> constraints
-val enforce_eq : universe constraint_function
-val enforce_leq : universe constraint_function
-val enforce_eq_level : universe_level constraint_function
-val enforce_leq_level : universe_level constraint_function
+val enforce_eq : Universe.t constraint_function
+val enforce_leq : Universe.t constraint_function
+val enforce_eq_level : Level.t constraint_function
+val enforce_leq_level : Level.t constraint_function
(** Type explanation is used to decorate error messages to provide
useful explanation why a given constraint is rejected. It is composed
@@ -196,17 +197,17 @@ val enforce_leq_level : universe_level constraint_function
system stores the graph and may result from combination of several
constraints...
*)
-type explanation = (constraint_type * universe) list
-type univ_inconsistency = constraint_type * universe * universe * explanation option
+type explanation = (constraint_type * Universe.t) list
+type univ_inconsistency = constraint_type * Universe.t * Universe.t * explanation option
exception UniverseInconsistency of univ_inconsistency
(** {6 Support for universe polymorphism } *)
(** Polymorphic maps from universe levels to 'a *)
-module LMap :
+module LMap :
sig
- include CMap.ExtS with type key = universe_level and module Set := LSet
+ include CMap.ExtS with type key = Level.t and module Set := LSet
val union : 'a t -> 'a t -> 'a t
(** [union x y] favors the bindings in the first map. *)
@@ -226,18 +227,18 @@ type 'a universe_map = 'a LMap.t
(** {6 Substitution} *)
-type universe_subst_fn = universe_level -> universe
-type universe_level_subst_fn = universe_level -> universe_level
+type universe_subst_fn = Level.t -> Universe.t
+type universe_level_subst_fn = Level.t -> Level.t
(** A full substitution, might involve algebraic universes *)
-type universe_subst = universe universe_map
-type universe_level_subst = universe_level universe_map
+type universe_subst = Universe.t universe_map
+type universe_level_subst = Level.t universe_map
val level_subst_of : universe_subst_fn -> universe_level_subst_fn
(** {6 Universe instances} *)
-module Instance :
+module Instance :
sig
type t
(** A universe instance represents a vector of argument universes
@@ -279,10 +280,11 @@ sig
end
type universe_instance = Instance.t
+[@@ocaml.deprecated "Use Instance.t"]
-val enforce_eq_instances : universe_instance constraint_function
+val enforce_eq_instances : Instance.t constraint_function
-type 'a puniverses = 'a * universe_instance
+type 'a puniverses = 'a * Instance.t
val out_punivs : 'a puniverses -> 'a
val in_punivs : 'a -> 'a puniverses
@@ -292,14 +294,14 @@ val eq_puniverses : ('a -> 'a -> bool) -> 'a puniverses -> 'a puniverses -> bool
representiong local universe variables and associated constraints *)
module UContext :
-sig
+sig
type t
val make : Instance.t constrained -> t
val empty : t
val is_empty : t -> bool
-
+
val instance : t -> Instance.t
val constraints : t -> constraints
@@ -314,9 +316,10 @@ sig
end
type universe_context = UContext.t
+[@@ocaml.deprecated "Use UContext.t"]
module AUContext :
-sig
+sig
type t
val repr : t -> UContext.t
@@ -340,6 +343,7 @@ sig
end
type abstract_universe_context = AUContext.t
+[@@ocaml.deprecated "Use AUContext.t"]
(** Universe info for inductive types: A context of universe levels
with universe constraints, representing local universe variables
@@ -354,49 +358,51 @@ module CumulativityInfo :
sig
type t
- val make : universe_context * universe_context -> t
+ val make : UContext.t * UContext.t -> t
val empty : t
val is_empty : t -> bool
- val univ_context : t -> universe_context
- val subtyp_context : t -> universe_context
+ val univ_context : t -> UContext.t
+ val subtyp_context : t -> UContext.t
(** This function takes a universe context representing constraints
of an inductive and a Instance.t of fresh universe names for the
subtyping (with the same length as the context in the given
universe context) and produces a UInfoInd.t that with the
trivial subtyping relation. *)
- val from_universe_context : universe_context -> universe_instance -> t
+ val from_universe_context : UContext.t -> Instance.t -> t
val subtyping_susbst : t -> universe_level_subst
end
type cumulativity_info = CumulativityInfo.t
+[@@ocaml.deprecated "Use CumulativityInfo.t"]
module ACumulativityInfo :
sig
type t
- val univ_context : t -> abstract_universe_context
- val subtyp_context : t -> abstract_universe_context
+ val univ_context : t -> AUContext.t
+ val subtyp_context : t -> AUContext.t
end
type abstract_cumulativity_info = ACumulativityInfo.t
+[@@ocaml.deprecated "Use ACumulativityInfo.t"]
(** Universe contexts (as sets) *)
module ContextSet :
-sig
- type t = universe_set constrained
+sig
+ type t = LSet.t constrained
val empty : t
val is_empty : t -> bool
- val singleton : universe_level -> t
+ val singleton : Level.t -> t
val of_instance : Instance.t -> t
- val of_set : universe_set -> t
+ val of_set : LSet.t -> t
val equal : t -> t -> bool
val union : t -> t -> t
@@ -406,39 +412,40 @@ sig
much smaller than the right one. *)
val diff : t -> t -> t
- val add_universe : universe_level -> t -> t
+ val add_universe : Level.t -> t -> t
val add_constraints : constraints -> t -> t
val add_instance : Instance.t -> t -> t
(** Arbitrary choice of linear order of the variables *)
val sort_levels : Level.t array -> Level.t array
- val to_context : t -> universe_context
- val of_context : universe_context -> t
+ val to_context : t -> UContext.t
+ val of_context : UContext.t -> t
val constraints : t -> constraints
- val levels : t -> universe_set
+ val levels : t -> LSet.t
end
(** A set of universes with universe constraints.
- We linearize the set to a list after typechecking.
+ We linearize the set to a list after typechecking.
Beware, representation could change.
*)
type universe_context_set = ContextSet.t
+[@@ocaml.deprecated "Use ContextSet.t"]
(** A value in a universe context (resp. context set). *)
-type 'a in_universe_context = 'a * universe_context
-type 'a in_universe_context_set = 'a * universe_context_set
+type 'a in_universe_context = 'a * UContext.t
+type 'a in_universe_context_set = 'a * ContextSet.t
val empty_level_subst : universe_level_subst
val is_empty_level_subst : universe_level_subst -> bool
(** Substitution of universes. *)
-val subst_univs_level_level : universe_level_subst -> universe_level -> universe_level
-val subst_univs_level_universe : universe_level_subst -> universe -> universe
+val subst_univs_level_level : universe_level_subst -> Level.t -> Level.t
+val subst_univs_level_universe : universe_level_subst -> Universe.t -> Universe.t
val subst_univs_level_constraints : universe_level_subst -> constraints -> constraints
val subst_univs_level_abstract_universe_context :
- universe_level_subst -> abstract_universe_context -> abstract_universe_context
-val subst_univs_level_instance : universe_level_subst -> universe_instance -> universe_instance
+ universe_level_subst -> AUContext.t -> AUContext.t
+val subst_univs_level_instance : universe_level_subst -> Instance.t -> Instance.t
(** Level to universe substitutions. *)
@@ -446,32 +453,32 @@ val empty_subst : universe_subst
val is_empty_subst : universe_subst -> bool
val make_subst : universe_subst -> universe_subst_fn
-val subst_univs_universe : universe_subst_fn -> universe -> universe
+val subst_univs_universe : universe_subst_fn -> Universe.t -> Universe.t
val subst_univs_constraints : universe_subst_fn -> constraints -> constraints
(** Substitution of instances *)
-val subst_instance_instance : universe_instance -> universe_instance -> universe_instance
-val subst_instance_universe : universe_instance -> universe -> universe
+val subst_instance_instance : Instance.t -> Instance.t -> Instance.t
+val subst_instance_universe : Instance.t -> Universe.t -> Universe.t
-val make_instance_subst : universe_instance -> universe_level_subst
-val make_inverse_instance_subst : universe_instance -> universe_level_subst
+val make_instance_subst : Instance.t -> universe_level_subst
+val make_inverse_instance_subst : Instance.t -> universe_level_subst
-val abstract_universes : universe_context -> universe_level_subst * abstract_universe_context
+val abstract_universes : UContext.t -> universe_level_subst * AUContext.t
-val abstract_cumulativity_info : cumulativity_info -> universe_level_subst * abstract_cumulativity_info
+val abstract_cumulativity_info : CumulativityInfo.t -> universe_level_subst * ACumulativityInfo.t
-val make_abstract_instance : abstract_universe_context -> universe_instance
+val make_abstract_instance : AUContext.t -> Instance.t
(** {6 Pretty-printing of universes. } *)
val pr_constraint_type : constraint_type -> Pp.t
val pr_constraints : (Level.t -> Pp.t) -> constraints -> Pp.t
-val pr_universe_context : (Level.t -> Pp.t) -> universe_context -> Pp.t
-val pr_cumulativity_info : (Level.t -> Pp.t) -> cumulativity_info -> Pp.t
-val pr_abstract_universe_context : (Level.t -> Pp.t) -> abstract_universe_context -> Pp.t
-val pr_abstract_cumulativity_info : (Level.t -> Pp.t) -> abstract_cumulativity_info -> Pp.t
-val pr_universe_context_set : (Level.t -> Pp.t) -> universe_context_set -> Pp.t
-val explain_universe_inconsistency : (Level.t -> Pp.t) ->
+val pr_universe_context : (Level.t -> Pp.t) -> UContext.t -> Pp.t
+val pr_cumulativity_info : (Level.t -> Pp.t) -> CumulativityInfo.t -> Pp.t
+val pr_abstract_universe_context : (Level.t -> Pp.t) -> AUContext.t -> Pp.t
+val pr_abstract_cumulativity_info : (Level.t -> Pp.t) -> ACumulativityInfo.t -> Pp.t
+val pr_universe_context_set : (Level.t -> Pp.t) -> ContextSet.t -> Pp.t
+val explain_universe_inconsistency : (Level.t -> Pp.t) ->
univ_inconsistency -> Pp.t
val pr_universe_level_subst : universe_level_subst -> Pp.t
@@ -479,23 +486,28 @@ val pr_universe_subst : universe_subst -> Pp.t
(** {6 Hash-consing } *)
-val hcons_univ : universe -> universe
+val hcons_univ : Universe.t -> Universe.t
val hcons_constraints : constraints -> constraints
-val hcons_universe_set : universe_set -> universe_set
-val hcons_universe_context : universe_context -> universe_context
-val hcons_abstract_universe_context : abstract_universe_context -> abstract_universe_context
-val hcons_universe_context_set : universe_context_set -> universe_context_set
-val hcons_cumulativity_info : cumulativity_info -> cumulativity_info
-val hcons_abstract_cumulativity_info : abstract_cumulativity_info -> abstract_cumulativity_info
+val hcons_universe_set : LSet.t -> LSet.t
+val hcons_universe_context : UContext.t -> UContext.t
+val hcons_abstract_universe_context : AUContext.t -> AUContext.t
+val hcons_universe_context_set : ContextSet.t -> ContextSet.t
+val hcons_cumulativity_info : CumulativityInfo.t -> CumulativityInfo.t
+val hcons_abstract_cumulativity_info : ACumulativityInfo.t -> ACumulativityInfo.t
(******)
(* deprecated: use qualified names instead *)
-val compare_levels : universe_level -> universe_level -> int
-val eq_levels : universe_level -> universe_level -> bool
+val compare_levels : Level.t -> Level.t -> int
+[@@ocaml.deprecated "Use Level.compare"]
+
+val eq_levels : Level.t -> Level.t -> bool
+[@@ocaml.deprecated "Use Level.equal"]
(** deprecated: Equality of formal universe expressions. *)
-val equal_universes : universe -> universe -> bool
+val equal_universes : Universe.t -> Universe.t -> bool
+[@@ocaml.deprecated "Use Universe.equal"]
(** Universes of constraints *)
-val universes_of_constraints : constraints -> universe_set
+val universes_of_constraints : constraints -> LSet.t
+[@@ocaml.deprecated "Use Constraint.universes_of"]
diff --git a/kernel/vars.mli b/kernel/vars.mli
index ae846fd42..964de4e95 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -141,8 +141,8 @@ val subst_univs_level_constr : universe_level_subst -> constr -> constr
val subst_univs_level_context : Univ.universe_level_subst -> Context.Rel.t -> Context.Rel.t
(** Instance substitution for polymorphism. *)
-val subst_instance_constr : universe_instance -> constr -> constr
-val subst_instance_context : universe_instance -> Context.Rel.t -> Context.Rel.t
+val subst_instance_constr : Instance.t -> constr -> constr
+val subst_instance_context : Instance.t -> Context.Rel.t -> Context.Rel.t
type id_key = Constant.t tableKey
val eq_id_key : id_key -> id_key -> bool
diff --git a/kernel/vconv.mli b/kernel/vconv.mli
index f4e680c69..7f727df47 100644
--- a/kernel/vconv.mli
+++ b/kernel/vconv.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Environ
open Reduction
diff --git a/kernel/vm.ml b/kernel/vm.ml
index 4e13881b8..51101f88e 100644
--- a/kernel/vm.ml
+++ b/kernel/vm.ml
@@ -7,7 +7,8 @@
(************************************************************************)
open Names
-open Term
+open Sorts
+open Constr
open Cbytecodes
external set_drawinstr : unit -> unit = "coq_set_drawinstr"
@@ -137,7 +138,7 @@ type vswitch = {
type atom =
| Aid of Vars.id_key
| Aind of inductive
- | Atype of Univ.universe
+ | Atype of Univ.Universe.t
(* Zippers *)
@@ -152,7 +153,7 @@ type stack = zipper list
type to_up = values
type whd =
- | Vsort of sorts
+ | Vsort of Sorts.t
| Vprod of vprod
| Vfun of vfun
| Vfix of vfix * arguments option
@@ -160,7 +161,7 @@ type whd =
| Vconstr_const of int
| Vconstr_block of vblock
| Vatom_stk of atom * stack
- | Vuniv_level of Univ.universe_level
+ | Vuniv_level of Univ.Level.t
(************************************************)
(* Abstract machine *****************************)
@@ -216,7 +217,7 @@ let apply_varray vf varray =
(* Destructors ***********************************)
(*************************************************)
-let uni_lvl_val (v : values) : Univ.universe_level =
+let uni_lvl_val (v : values) : Univ.Level.t =
let whd = Obj.magic v in
match whd with
| Vuniv_level lvl -> lvl
diff --git a/kernel/vm.mli b/kernel/vm.mli
index ea38ee43f..bc38452d4 100644
--- a/kernel/vm.mli
+++ b/kernel/vm.mli
@@ -1,5 +1,13 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
open Names
-open Term
+open Constr
open Cbytecodes
(** Debug printing *)
@@ -23,7 +31,7 @@ type arguments
type atom =
| Aid of Vars.id_key
| Aind of inductive
- | Atype of Univ.universe
+ | Atype of Univ.Universe.t
(** Zippers *)
@@ -38,7 +46,7 @@ type stack = zipper list
type to_up
type whd =
- | Vsort of sorts
+ | Vsort of Sorts.t
| Vprod of vprod
| Vfun of vfun
| Vfix of vfix * arguments option
@@ -46,7 +54,7 @@ type whd =
| Vconstr_const of int
| Vconstr_block of vblock
| Vatom_stk of atom * stack
- | Vuniv_level of Univ.universe_level
+ | Vuniv_level of Univ.Level.t
(** For debugging purposes only *)
@@ -66,7 +74,7 @@ external val_of_annot_switch : annot_switch -> values = "%identity"
(** Destructors *)
val whd_val : values -> whd
-val uni_lvl_val : values -> Univ.universe_level
+val uni_lvl_val : values -> Univ.Level.t
(** Arguments *)
diff --git a/library/declaremods.mli b/library/declaremods.mli
index 90e86cd02..42e5f4b13 100644
--- a/library/declaremods.mli
+++ b/library/declaremods.mli
@@ -72,7 +72,7 @@ type library_objects
val register_library :
library_name ->
Safe_typing.compiled_library -> library_objects -> Safe_typing.vodigest ->
- Univ.universe_context_set -> unit
+ Univ.ContextSet.t -> unit
val get_library_native_symbols : library_name -> Nativecode.symbols
diff --git a/library/decls.ml b/library/decls.ml
index 973fe144d..a4259f6ca 100644
--- a/library/decls.ml
+++ b/library/decls.ml
@@ -19,7 +19,7 @@ module NamedDecl = Context.Named.Declaration
(** Datas associated to section variables and local definitions *)
type variable_data =
- DirPath.t * bool (* opacity *) * Univ.universe_context_set * polymorphic * logical_kind
+ DirPath.t * bool (* opacity *) * Univ.ContextSet.t * polymorphic * logical_kind
let vartab =
Summary.ref (Id.Map.empty : variable_data Id.Map.t) ~name:"VARIABLE"
diff --git a/library/decls.mli b/library/decls.mli
index 524377257..1b7f137a4 100644
--- a/library/decls.mli
+++ b/library/decls.mli
@@ -17,14 +17,14 @@ open Decl_kinds
(** Registration and access to the table of variable *)
type variable_data =
- DirPath.t * bool (** opacity *) * Univ.universe_context_set * polymorphic * logical_kind
+ DirPath.t * bool (** opacity *) * Univ.ContextSet.t * polymorphic * logical_kind
val add_variable_data : variable -> variable_data -> unit
val variable_path : variable -> DirPath.t
val variable_secpath : variable -> qualid
val variable_kind : variable -> logical_kind
val variable_opacity : variable -> bool
-val variable_context : variable -> Univ.universe_context_set
+val variable_context : variable -> Univ.ContextSet.t
val variable_polymorphic : variable -> polymorphic
val variable_exists : variable -> bool
diff --git a/library/global.ml b/library/global.ml
index 28b9e66f8..43097dc5d 100644
--- a/library/global.ml
+++ b/library/global.ml
@@ -233,7 +233,7 @@ let universes_of_global gr =
(** Global universe names *)
type universe_names =
- (polymorphic * Univ.universe_level) Id.Map.t * Id.t Univ.LMap.t
+ (polymorphic * Univ.Level.t) Id.Map.t * Id.t Univ.LMap.t
let global_universes =
Summary.ref ~name:"Global universe names"
diff --git a/library/global.mli b/library/global.mli
index 1aff0a376..51fe53181 100644
--- a/library/global.mli
+++ b/library/global.mli
@@ -32,7 +32,7 @@ val set_typing_flags : Declarations.typing_flags -> unit
(** Variables, Local definitions, constants, inductive types *)
val push_named_assum : (Id.t * Constr.types * bool) Univ.in_universe_context_set -> unit
-val push_named_def : (Id.t * Safe_typing.private_constants Entries.definition_entry) -> Univ.universe_context_set
+val push_named_def : (Id.t * Safe_typing.private_constants Entries.definition_entry) -> Univ.ContextSet.t
val export_private_constants : in_section:bool ->
Safe_typing.private_constants Entries.constant_entry ->
@@ -46,8 +46,8 @@ val add_mind :
(** Extra universe constraints *)
val add_constraints : Univ.constraints -> unit
-val push_context : bool -> Univ.universe_context -> unit
-val push_context_set : bool -> Univ.universe_context_set -> unit
+val push_context : bool -> Univ.UContext.t -> unit
+val push_context_set : bool -> Univ.ContextSet.t -> unit
(** Non-interactive modules and module types *)
@@ -93,18 +93,18 @@ val mind_of_delta_kn : KerName.t -> MutInd.t
val opaque_tables : unit -> Opaqueproof.opaquetab
-val body_of_constant : Constant.t -> (Term.constr * Univ.AUContext.t) option
+val body_of_constant : Constant.t -> (Constr.constr * Univ.AUContext.t) option
(** Returns the body of the constant if it has any, and the polymorphic context
it lives in. For monomorphic constant, the latter is empty, and for
polymorphic constants, the term contains De Bruijn universe variables that
need to be instantiated. *)
-val body_of_constant_body : Declarations.constant_body -> (Term.constr * Univ.AUContext.t) option
+val body_of_constant_body : Declarations.constant_body -> (Constr.constr * Univ.AUContext.t) option
(** Same as {!body_of_constant} but on {!Declarations.constant_body}. *)
(** Global universe name <-> level mapping *)
type universe_names =
- (Decl_kinds.polymorphic * Univ.universe_level) Id.Map.t * Id.t Univ.LMap.t
+ (Decl_kinds.polymorphic * Univ.Level.t) Id.Map.t * Id.t Univ.LMap.t
val global_universe_names : unit -> universe_names
val set_global_universe_names : universe_names -> unit
@@ -115,7 +115,7 @@ val start_library : DirPath.t -> ModPath.t
val export : ?except:Future.UUIDSet.t -> DirPath.t ->
ModPath.t * Safe_typing.compiled_library * Safe_typing.native_library
val import :
- Safe_typing.compiled_library -> Univ.universe_context_set -> Safe_typing.vodigest ->
+ Safe_typing.compiled_library -> Univ.ContextSet.t -> Safe_typing.vodigest ->
ModPath.t
(** {6 Misc } *)
@@ -147,12 +147,12 @@ val type_of_global_in_context : Environ.env ->
constants, it does not matter. *)
(** Returns the universe context of the global reference (whatever its polymorphic status is). *)
-val universes_of_global : Globnames.global_reference -> Univ.abstract_universe_context
+val universes_of_global : Globnames.global_reference -> Univ.AUContext.t
(** {6 Retroknowledge } *)
val register :
- Retroknowledge.field -> Term.constr -> Term.constr -> unit
+ Retroknowledge.field -> Constr.constr -> Constr.constr -> unit
val register_inline : Constant.t -> unit
diff --git a/library/globnames.ml b/library/globnames.ml
index 9f652896c..9d7ab2db8 100644
--- a/library/globnames.ml
+++ b/library/globnames.ml
@@ -8,7 +8,7 @@
open CErrors
open Names
-open Term
+open Constr
open Mod_subst
open Libnames
@@ -72,7 +72,7 @@ let canonical_gr = function
| ConstructRef ((kn,i),j )-> ConstructRef((MutInd.make1(MutInd.canonical kn),i),j)
| VarRef id -> VarRef id
-let global_of_constr c = match kind_of_term c with
+let global_of_constr c = match kind c with
| Const (sp,u) -> ConstRef sp
| Ind (ind_sp,u) -> IndRef ind_sp
| Construct (cstr_cp,u) -> ConstructRef cstr_cp
@@ -80,7 +80,7 @@ let global_of_constr c = match kind_of_term c with
| _ -> raise Not_found
let is_global c t =
- match c, kind_of_term t with
+ match c, kind t with
| ConstRef c, Const (c', _) -> Constant.equal c c'
| IndRef i, Ind (i', _) -> eq_ind i i'
| ConstructRef i, Construct (i', _) -> eq_constructor i i'
diff --git a/library/globnames.mli b/library/globnames.mli
index 361631bd3..5c484b391 100644
--- a/library/globnames.mli
+++ b/library/globnames.mli
@@ -8,7 +8,7 @@
open Util
open Names
-open Term
+open Constr
open Mod_subst
(** {6 Global reference is a kernel side type for all references together } *)
diff --git a/library/heads.ml b/library/heads.ml
index d2cfc0990..8b8e407f7 100644
--- a/library/heads.ml
+++ b/library/heads.ml
@@ -9,6 +9,7 @@
open Util
open Names
open Term
+open Constr
open Vars
open Mod_subst
open Environ
@@ -57,7 +58,7 @@ let variable_head id = Evalrefmap.find (EvalVarRef id) !head_map
let constant_head cst = Evalrefmap.find (EvalConstRef cst) !head_map
let kind_of_head env t =
- let rec aux k l t b = match kind_of_term (Reduction.whd_betaiotazeta env t) with
+ let rec aux k l t b = match kind (Reduction.whd_betaiotazeta env t) with
| Rel n when n > k -> NotImmediatelyComputableHead
| Rel n -> FlexibleHead (k,k+1-n,List.length l,b)
| Var id ->
diff --git a/library/heads.mli b/library/heads.mli
index 1ce66c841..8ad5c0f14 100644
--- a/library/heads.mli
+++ b/library/heads.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
(** This module is about the computation of an approximation of the
diff --git a/library/lib.ml b/library/lib.ml
index a6fa27be8..df9563e45 100644
--- a/library/lib.ml
+++ b/library/lib.ml
@@ -417,8 +417,8 @@ type abstr_list = abstr_info Names.Cmap.t * abstr_info Names.Mindmap.t
type secentry =
| Variable of (Names.Id.t * Decl_kinds.binding_kind *
- Decl_kinds.polymorphic * Univ.universe_context_set)
- | Context of Univ.universe_context_set
+ Decl_kinds.polymorphic * Univ.ContextSet.t)
+ | Context of Univ.ContextSet.t
let sectab =
Summary.ref ([] : (secentry list * Opaqueproof.work_list * abstr_list) list)
diff --git a/library/lib.mli b/library/lib.mli
index f941fd6d4..721e2896f 100644
--- a/library/lib.mli
+++ b/library/lib.mli
@@ -162,11 +162,11 @@ val section_segment_of_constant : Names.Constant.t -> abstr_info
val section_segment_of_mutual_inductive: Names.MutInd.t -> abstr_info
val variable_section_segment_of_reference : Globnames.global_reference -> variable_context
-val section_instance : Globnames.global_reference -> Univ.universe_instance * Names.Id.t array
+val section_instance : Globnames.global_reference -> Univ.Instance.t * Names.Id.t array
val is_in_section : Globnames.global_reference -> bool
-val add_section_variable : Names.Id.t -> Decl_kinds.binding_kind -> Decl_kinds.polymorphic -> Univ.universe_context_set -> unit
-val add_section_context : Univ.universe_context_set -> unit
+val add_section_variable : Names.Id.t -> Decl_kinds.binding_kind -> Decl_kinds.polymorphic -> Univ.ContextSet.t -> unit
+val add_section_context : Univ.ContextSet.t -> unit
val add_section_constant : Decl_kinds.polymorphic ->
Names.Constant.t -> Context.Named.t -> unit
val add_section_kn : Decl_kinds.polymorphic ->
diff --git a/library/library.ml b/library/library.ml
index e2832ecdc..840fe563a 100644
--- a/library/library.ml
+++ b/library/library.ml
@@ -97,7 +97,7 @@ type library_t = {
library_deps : (compilation_unit_name * Safe_typing.vodigest) array;
library_imports : compilation_unit_name array;
library_digests : Safe_typing.vodigest;
- library_extra_univs : Univ.universe_context_set;
+ library_extra_univs : Univ.ContextSet.t;
}
type library_summary = {
@@ -360,9 +360,9 @@ type 'a table_status =
| Fetched of 'a Future.computation array
let opaque_tables =
- ref (LibraryMap.empty : (Term.constr table_status) LibraryMap.t)
+ ref (LibraryMap.empty : (Constr.constr table_status) LibraryMap.t)
let univ_tables =
- ref (LibraryMap.empty : (Univ.universe_context_set table_status) LibraryMap.t)
+ ref (LibraryMap.empty : (Univ.ContextSet.t table_status) LibraryMap.t)
let add_opaque_table dp st =
opaque_tables := LibraryMap.add dp st !opaque_tables
@@ -408,9 +408,9 @@ let () =
type seg_sum = summary_disk
type seg_lib = library_disk
type seg_univ = (* true = vivo, false = vi *)
- Univ.universe_context_set Future.computation array * Univ.universe_context_set * bool
+ Univ.ContextSet.t Future.computation array * Univ.ContextSet.t * bool
type seg_discharge = Opaqueproof.cooking_info list array
-type seg_proofs = Term.constr Future.computation array
+type seg_proofs = Constr.constr Future.computation array
let mk_library sd md digests univs =
{
diff --git a/library/library.mli b/library/library.mli
index 6c624ce52..63e7b95bb 100644
--- a/library/library.mli
+++ b/library/library.mli
@@ -29,9 +29,9 @@ val require_library_from_dirpath : (DirPath.t * string) list -> bool option -> u
type seg_sum
type seg_lib
type seg_univ = (* cst, all_cst, finished? *)
- Univ.universe_context_set Future.computation array * Univ.universe_context_set * bool
+ Univ.ContextSet.t Future.computation array * Univ.ContextSet.t * bool
type seg_discharge = Opaqueproof.cooking_info list array
-type seg_proofs = Term.constr Future.computation array
+type seg_proofs = Constr.constr Future.computation array
(** Open a module (or a library); if the boolean is true then it's also
an export otherwise just a simple import *)
diff --git a/library/nameops.mli b/library/nameops.mli
index 58cd6ed4e..26f300b61 100644
--- a/library/nameops.mli
+++ b/library/nameops.mli
@@ -131,5 +131,5 @@ val coq_string : string (** "Coq" *)
val default_root_prefix : DirPath.t
(** Metavariables *)
-val pr_meta : Term.metavariable -> Pp.t
-val string_of_meta : Term.metavariable -> string
+val pr_meta : Constr.metavariable -> Pp.t
+val string_of_meta : Constr.metavariable -> string
diff --git a/library/univops.ml b/library/univops.ml
index 3bafb824d..9dc138eb8 100644
--- a/library/univops.ml
+++ b/library/univops.ml
@@ -6,18 +6,18 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Univ
let universes_of_constr c =
let rec aux s c =
- match kind_of_term c with
+ match kind c with
| Const (_, u) | Ind (_, u) | Construct (_, u) ->
LSet.fold LSet.add (Instance.levels u) s
| Sort u when not (Sorts.is_small u) ->
- let u = univ_of_sort u in
+ let u = Term.univ_of_sort u in
LSet.fold LSet.add (Universe.levels u) s
- | _ -> fold_constr aux s c
+ | _ -> Constr.fold aux s c
in aux LSet.empty c
let restrict_universe_context (univs,csts) s =
diff --git a/library/univops.mli b/library/univops.mli
index 09147cb41..9af568bcb 100644
--- a/library/univops.mli
+++ b/library/univops.mli
@@ -6,10 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Univ
(** Shrink a universe context to a restricted set of variables *)
-val universes_of_constr : constr -> universe_set
-val restrict_universe_context : universe_context_set -> universe_set -> universe_context_set
+val universes_of_constr : constr -> LSet.t
+val restrict_universe_context : ContextSet.t -> LSet.t -> ContextSet.t
diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml
index 6281b2675..da8955f0d 100644
--- a/plugins/btauto/refl_btauto.ml
+++ b/plugins/btauto/refl_btauto.ml
@@ -12,12 +12,12 @@ let get_inductive dir s =
let glob_ref () = Coqlib.find_reference contrib_name ("Coq" :: dir) s in
Lazy.from_fun (fun () -> Globnames.destIndRef (glob_ref ()))
-let decomp_term sigma (c : Term.constr) =
- Term.kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast sigma (EConstr.of_constr c)))
+let decomp_term sigma (c : Constr.t) =
+ Constr.kind (EConstr.Unsafe.to_constr (Termops.strip_outer_cast sigma (EConstr.of_constr c)))
-let lapp c v = Term.mkApp (Lazy.force c, v)
+let lapp c v = Constr.mkApp (Lazy.force c, v)
-let (===) = Term.eq_constr
+let (===) = Constr.equal
module CoqList = struct
let path = ["Init"; "Datatypes"]
@@ -53,17 +53,11 @@ end
module Env = struct
- module ConstrHashed = struct
- type t = Term.constr
- let equal = Term.eq_constr
- let hash = Term.hash_constr
- end
-
- module ConstrHashtbl = Hashtbl.Make (ConstrHashed)
+ module ConstrHashtbl = Hashtbl.Make (Constr)
type t = (int ConstrHashtbl.t * int ref)
- let add (tbl, off) (t : Term.constr) =
+ let add (tbl, off) (t : Constr.t) =
try ConstrHashtbl.find tbl t
with
| Not_found ->
@@ -103,7 +97,7 @@ module Bool = struct
| Negb of t
| Ifb of t * t * t
- let quote (env : Env.t) sigma (c : Term.constr) : t =
+ let quote (env : Env.t) sigma (c : Constr.t) : t =
let trueb = Lazy.force trueb in
let falseb = Lazy.force falseb in
let andb = Lazy.force andb in
@@ -170,7 +164,7 @@ module Btauto = struct
| Bool.Xorb (b1, b2) -> lapp f_xor [|convert b1; convert b2|]
| Bool.Ifb (b1, b2, b3) -> lapp f_ifb [|convert b1; convert b2; convert b3|]
- let convert_env env : Term.constr =
+ let convert_env env : Constr.t =
CoqList.of_list (Lazy.force Bool.typ) env
let reify env t = lapp eval [|convert_env env; convert t|]
@@ -249,7 +243,7 @@ module Btauto = struct
let env = Env.to_list env in
let fl = reify env fl in
let fr = reify env fr in
- let changed_gl = Term.mkApp (c, [|typ; fl; fr|]) in
+ let changed_gl = Constr.mkApp (c, [|typ; fl; fr|]) in
let changed_gl = EConstr.of_constr changed_gl in
Tacticals.New.tclTHENLIST [
Tactics.change_concl changed_gl;
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 182821322..faabd7c14 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -11,14 +11,15 @@
(* Plus some e-matching and constructor handling by P. Corbineau *)
open CErrors
-open Util
open Pp
open Goptions
open Names
open Term
+open Constr
open Vars
open Tacmach
open Evd
+open Util
let init_size=5
@@ -154,7 +155,7 @@ let rec term_equal t1 t2 =
open Hashset.Combine
let rec hash_term = function
- | Symb c -> combine 1 (hash_constr c)
+ | Symb c -> combine 1 (Constr.hash c)
| Product (s1, s2) -> combine3 2 (Sorts.hash s1) (Sorts.hash s2)
| Eps i -> combine 3 (Id.hash i)
| Appli (t1, t2) -> combine3 4 (hash_term t1) (hash_term t2)
@@ -215,7 +216,7 @@ type representative=
mutable lfathers:Int.Set.t;
mutable fathers:Int.Set.t;
mutable inductive_status: inductive_status;
- class_type : Term.types;
+ class_type : types;
mutable functions: Int.Set.t PafMap.t} (*pac -> term = app(constr,t) *)
type cl = Rep of representative| Eqto of int*equality
@@ -232,7 +233,7 @@ type node =
module Constrhash = Hashtbl.Make
(struct type t = constr
let equal = eq_constr_nounivs
- let hash = hash_constr
+ let hash = Constr.hash
end)
module Typehash = Constrhash
@@ -438,7 +439,7 @@ and applist_proj c l =
| Symb s -> applist_projection s l
| _ -> applistc (constr_of_term c) l
and applist_projection c l =
- match kind_of_term c with
+ match Constr.kind c with
| Const c when Environ.is_projection (fst c) (Global.env()) ->
let p = Projection.make (fst c) false in
(match l with
@@ -454,7 +455,7 @@ and applist_projection c l =
let rec canonize_name sigma c =
let c = EConstr.Unsafe.to_constr c in
let func c = canonize_name sigma (EConstr.of_constr c) in
- match kind_of_term c with
+ match Constr.kind c with
| Const (kn,u) ->
let canon_const = Constant.make1 (Constant.canonical kn) in
(mkConstU (canon_const,u))
diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli
index f904aa3e6..23cd2161d 100644
--- a/plugins/cc/ccalgo.mli
+++ b/plugins/cc/ccalgo.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Util
-open Term
+open Constr
open Names
type pa_constructor =
@@ -85,7 +85,7 @@ type representative=
mutable lfathers:Int.Set.t;
mutable fathers:Int.Set.t;
mutable inductive_status: inductive_status;
- class_type : Term.types;
+ class_type : types;
mutable functions: Int.Set.t PafMap.t} (*pac -> term = app(constr,t) *)
type cl = Rep of representative| Eqto of int*equality
diff --git a/plugins/cc/ccproof.ml b/plugins/cc/ccproof.ml
index a43a167e8..97efaced8 100644
--- a/plugins/cc/ccproof.ml
+++ b/plugins/cc/ccproof.ml
@@ -10,7 +10,7 @@
(* proof-trees that will be transformed into proof-terms in cctac.ml4 *)
open CErrors
-open Term
+open Constr
open Ccalgo
open Pp
diff --git a/plugins/cc/ccproof.mli b/plugins/cc/ccproof.mli
index 9f53123db..a3e450134 100644
--- a/plugins/cc/ccproof.mli
+++ b/plugins/cc/ccproof.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Ccalgo
-open Term
+open Constr
type rule=
Ax of constr
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 150319f6b..7dec34d4d 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -13,6 +13,7 @@ open Names
open Inductiveops
open Declarations
open Term
+open Constr
open EConstr
open Vars
open Tactics
@@ -76,11 +77,11 @@ let rec decompose_term env sigma t=
let (mind,i_ind),u = c in
let u = EInstance.kind sigma u in
let canon_mind = MutInd.make1 (MutInd.canonical mind) in
- let canon_ind = canon_mind,i_ind in (Symb (Term.mkIndU (canon_ind,u)))
+ let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u)))
| Const (c,u) ->
let u = EInstance.kind sigma u in
let canon_const = Constant.make1 (Constant.canonical c) in
- (Symb (Term.mkConstU (canon_const,u)))
+ (Symb (Constr.mkConstU (canon_const,u)))
| Proj (p, c) ->
let canon_const kn = Constant.make1 (Constant.canonical kn) in
let p' = Projection.map canon_const p in
@@ -198,7 +199,7 @@ let make_prb gls depth additionnal_terms =
(fun decl ->
let id = NamedDecl.get_id decl in
begin
- let cid=Term.mkVar id in
+ let cid=Constr.mkVar id in
match litteral_of_constr env sigma (NamedDecl.get_type decl) with
`Eq (t,a,b) -> add_equality state cid a b
| `Neq (t,a,b) -> add_disequality state (Hyp cid) a b
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index 6d3d4b743..fb65a8639 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -6,9 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Constr
open Context.Named.Declaration
-let map_const_entry_body (f:Term.constr->Term.constr) (x:Safe_typing.private_constants Entries.const_entry_body)
+let map_const_entry_body (f:constr->constr) (x:Safe_typing.private_constants Entries.const_entry_body)
: Safe_typing.private_constants Entries.const_entry_body =
Future.chain x begin fun ((b,ctx),fx) ->
(f b , ctx) , fx
@@ -67,7 +68,7 @@ let start_deriving f suchthat lemma =
let f_def = { f_def with Entries.const_entry_opaque = false } in
let f_def = Entries.DefinitionEntry f_def , Decl_kinds.(IsDefinition Definition) in
let f_kn = Declare.declare_constant f f_def in
- let f_kn_term = Term.mkConst f_kn in
+ let f_kn_term = mkConst f_kn in
(** In the type and body of the proof of [suchthat] there can be
references to the variable [f]. It needs to be replaced by
references to the constant [f] declared above. This substitution
diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml
index 3c46d5c43..bc84df76b 100644
--- a/plugins/extraction/extract_env.ml
+++ b/plugins/extraction/extract_env.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Miniml
-open Term
+open Constr
open Declarations
open Names
open ModPath
@@ -138,7 +138,7 @@ let check_arity env cb =
let check_fix env cb i =
match cb.const_body with
| Def lbody ->
- (match kind_of_term (Mod_subst.force_constr lbody) with
+ (match Constr.kind (Mod_subst.force_constr lbody) with
| Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd)
| CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd)
| _ -> raise Impossible)
@@ -146,8 +146,8 @@ let check_fix env cb i =
let prec_declaration_equal (na1, ca1, ta1) (na2, ca2, ta2) =
Array.equal Name.equal na1 na2 &&
- Array.equal eq_constr ca1 ca2 &&
- Array.equal eq_constr ta1 ta2
+ Array.equal Constr.equal ca1 ca2 &&
+ Array.equal Constr.equal ta1 ta2
let factor_fix env l cb msb =
let _,recd as check = check_fix env cb 0 in
diff --git a/plugins/extraction/extract_env.mli b/plugins/extraction/extract_env.mli
index 7bbb825b1..dd8617738 100644
--- a/plugins/extraction/extract_env.mli
+++ b/plugins/extraction/extract_env.mli
@@ -34,4 +34,4 @@ val print_one_decl :
(* Used by Extraction Compute *)
val structure_for_compute :
- Term.constr -> (Miniml.ml_decl list) * Miniml.ml_ast * Miniml.ml_type
+ Constr.t -> (Miniml.ml_decl list) * Miniml.ml_ast * Miniml.ml_type
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index a227478d0..47e812319 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -10,6 +10,7 @@
open Util
open Names
open Term
+open Constr
open Vars
open Declarations
open Declareops
@@ -81,7 +82,7 @@ let whd_betaiotazeta t =
let rec flag_of_type env t : flag =
let t = whd_all env t in
- match kind_of_term t with
+ match Constr.kind t with
| Prod (x,t,c) -> flag_of_type (push_rel (LocalAssum (x,t)) env) c
| Sort s when Sorts.is_prop s -> (Logic,TypeScheme)
| Sort _ -> (Info,TypeScheme)
@@ -111,14 +112,14 @@ let push_rel_assum (n, t) env =
(*s [type_sign] gernerates a signature aimed at treating a type application. *)
let rec type_sign env c =
- match kind_of_term (whd_all env c) with
+ match Constr.kind (whd_all env c) with
| Prod (n,t,d) ->
(if is_info_scheme env t then Keep else Kill Kprop)
:: (type_sign (push_rel_assum (n,t) env) d)
| _ -> []
let rec type_scheme_nb_args env c =
- match kind_of_term (whd_all env c) with
+ match Constr.kind (whd_all env c) with
| Prod (n,t,d) ->
let n = type_scheme_nb_args (push_rel_assum (n,t) env) d in
if is_info_scheme env t then n+1 else n
@@ -145,7 +146,7 @@ let make_typvar n vl =
next_ident_away id' vl
let rec type_sign_vl env c =
- match kind_of_term (whd_all env c) with
+ match Constr.kind (whd_all env c) with
| Prod (n,t,d) ->
let s,vl = type_sign_vl (push_rel_assum (n,t) env) d in
if not (is_info_scheme env t) then Kill Kprop::s, vl
@@ -153,7 +154,7 @@ let rec type_sign_vl env c =
| _ -> [],[]
let rec nb_default_params env c =
- match kind_of_term (whd_all env c) with
+ match Constr.kind (whd_all env c) with
| Prod (n,t,d) ->
let n = nb_default_params (push_rel_assum (n,t) env) d in
if is_default env t then n+1 else n
@@ -207,7 +208,7 @@ let parse_ind_args si args relmax =
| [] -> Int.Map.empty
| Kill _ :: s -> parse (i+1) j s
| Keep :: s ->
- (match kind_of_term args.(i-1) with
+ (match Constr.kind args.(i-1) with
| Rel k -> Int.Map.add (relmax+1-k) j (parse (i+1) (j+1) s)
| _ -> parse (i+1) (j+1) s)
in parse 1 1 si
@@ -224,7 +225,7 @@ let parse_ind_args si args relmax =
let rec extract_type env db j c args =
- match kind_of_term (whd_betaiotazeta c) with
+ match Constr.kind (whd_betaiotazeta c) with
| App (d, args') ->
(* We just accumulate the arguments. *)
extract_type env db j d (Array.to_list args' @ args)
@@ -299,7 +300,7 @@ let rec extract_type env db j c args =
| Proj (p,t) ->
(* Let's try to reduce, if it hasn't already been done. *)
if Projection.unfolded p then Tunknown
- else extract_type env db j (Term.mkProj (Projection.unfold p, t)) args
+ else extract_type env db j (mkProj (Projection.unfold p, t)) args
| Case _ | Fix _ | CoFix _ -> Tunknown
| _ -> assert false
@@ -331,7 +332,7 @@ and extract_type_scheme env db c p =
if Int.equal p 0 then extract_type env db 0 c []
else
let c = whd_betaiotazeta c in
- match kind_of_term c with
+ match Constr.kind c with
| Lambda (n,t,d) ->
extract_type_scheme (push_rel_assum (n,t) env) db d (p-1)
| _ ->
@@ -415,8 +416,8 @@ and extract_really_ind env kn mib =
let t = snd (decompose_prod_n npar types.(j)) in
let prods,head = dest_prod epar t in
let nprods = List.length prods in
- let args = match kind_of_term head with
- | App (f,args) -> args (* [kind_of_term f = Ind ip] *)
+ let args = match Constr.kind head with
+ | App (f,args) -> args (* [Constr.kind f = Ind ip] *)
| _ -> [||]
in
let dbmap = parse_ind_args p.ip_sign args (nprods + npar) in
@@ -444,7 +445,7 @@ and extract_really_ind env kn mib =
if Option.is_empty mib.mind_record then raise (I Standard);
(* Now we're sure it's a record. *)
(* First, we find its field names. *)
- let rec names_prod t = match kind_of_term t with
+ let rec names_prod t = match Constr.kind t with
| Prod(n,_,t) -> n::(names_prod t)
| LetIn(_,_,_,t) -> names_prod t
| Cast(t,_,_) -> names_prod t
@@ -503,7 +504,7 @@ and extract_really_ind env kn mib =
*)
and extract_type_cons env db dbmap c i =
- match kind_of_term (whd_all env c) with
+ match Constr.kind (whd_all env c) with
| Prod (n,t,d) ->
let env' = push_rel_assum (n,t) env in
let db' = (try Int.Map.find i dbmap with Not_found -> 0) :: db in
@@ -564,7 +565,7 @@ let record_constant_type env kn opt_typ =
(* [mlt] is the ML type we want our extraction of [(c args)] to have. *)
let rec extract_term env mle mlt c args =
- match kind_of_term c with
+ match Constr.kind c with
| App (f,a) ->
extract_term env mle mlt f (Array.to_list a @ args)
| Lambda (n, t, d) ->
@@ -874,7 +875,7 @@ let decomp_lams_eta_n n m env c t =
(* Let's try to identify some situation where extracted code
will allow generalisation of type variables *)
-let rec gentypvar_ok c = match kind_of_term c with
+let rec gentypvar_ok c = match Constr.kind c with
| Lambda _ | Const _ -> true
| App (c,v) ->
(* if all arguments are variables, these variables will
diff --git a/plugins/extraction/extraction.mli b/plugins/extraction/extraction.mli
index e1d43f340..b15b88ed2 100644
--- a/plugins/extraction/extraction.mli
+++ b/plugins/extraction/extraction.mli
@@ -9,7 +9,7 @@
(*s Extraction from Coq terms to Miniml. *)
open Names
-open Term
+open Constr
open Declarations
open Environ
open Miniml
diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli
index cc93f294b..e52e419fd 100644
--- a/plugins/extraction/table.mli
+++ b/plugins/extraction/table.mli
@@ -180,7 +180,7 @@ val implicits_of_global : global_reference -> Int.Set.t
(*s Table for user-given custom ML extractions. *)
(* UGLY HACK: registration of a function defined in [extraction.ml] *)
-val type_scheme_nb_args_hook : (Environ.env -> Term.constr -> int) Hook.t
+val type_scheme_nb_args_hook : (Environ.env -> Constr.t -> int) Hook.t
val is_custom : global_reference -> bool
val is_inline_custom : global_reference -> bool
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index db1a46a03..c55040df0 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -8,7 +8,7 @@
open Hipattern
open Names
-open Term
+open Constr
open EConstr
open Vars
open Termops
@@ -39,7 +39,7 @@ exception Is_atom of constr
let meta_succ m = m+1
let rec nb_prod_after n c=
- match kind_of_term c with
+ match Constr.kind c with
| Prod (_,_,b) ->if n>0 then nb_prod_after (n-1) b else
1+(nb_prod_after 0 b)
| _ -> 0
diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli
index 106c469c6..3b6b711c0 100644
--- a/plugins/firstorder/formula.mli
+++ b/plugins/firstorder/formula.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open EConstr
open Globnames
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index c2606dbe8..3409471a7 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -24,7 +24,7 @@ open Misctypes
open Context.Rel.Declaration
let compare_instance inst1 inst2=
- let cmp c1 c2 = OrderedConstr.compare (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in
+ let cmp c1 c2 = Constr.compare (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in
match inst1,inst2 with
Phantom(d1),Phantom(d2)->
(cmp d1 d2)
diff --git a/plugins/firstorder/rules.mli b/plugins/firstorder/rules.mli
index d8d4c1a38..5c46f4cec 100644
--- a/plugins/firstorder/rules.mli
+++ b/plugins/firstorder/rules.mli
@@ -6,9 +6,9 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
-open EConstr
open Names
+open Constr
+open EConstr
open Globnames
type tactic = unit Proofview.tactic
diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml
index 05194164b..ea2d076ed 100644
--- a/plugins/firstorder/sequent.ml
+++ b/plugins/firstorder/sequent.ml
@@ -54,13 +54,7 @@ struct
(priority e1.pat) - (priority e2.pat)
end
-module OrderedConstr=
-struct
- type t=Term.constr
- let compare=Term.compare
-end
-
-type h_item = global_reference * (int*Term.constr) option
+type h_item = global_reference * (int*Constr.t) option
module Hitem=
struct
@@ -70,13 +64,13 @@ struct
if c = 0 then
let cmp (i1, c1) (i2, c2) =
let c = Int.compare i1 i2 in
- if c = 0 then OrderedConstr.compare c1 c2 else c
+ if c = 0 then Constr.compare c1 c2 else c
in
Option.compare cmp co1 co2
else c
end
-module CM=Map.Make(OrderedConstr)
+module CM=Map.Make(Constr)
module History=Set.Make(Hitem)
diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli
index ca6079c8b..7f4a6dd86 100644
--- a/plugins/firstorder/sequent.mli
+++ b/plugins/firstorder/sequent.mli
@@ -10,11 +10,9 @@ open EConstr
open Formula
open Globnames
-module OrderedConstr: Set.OrderedType with type t=Term.constr
+module CM: CSig.MapS with type key=Constr.t
-module CM: CSig.MapS with type key=Term.constr
-
-type h_item = global_reference * (int*Term.constr) option
+type h_item = global_reference * (int*Constr.t) option
module History: Set.S with type elt = h_item
diff --git a/plugins/firstorder/unify.mli b/plugins/firstorder/unify.mli
index d3e8aeee8..390aa8c85 100644
--- a/plugins/firstorder/unify.mli
+++ b/plugins/firstorder/unify.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open EConstr
exception UFAIL of constr*constr
diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml
index 68af1b3b6..d9e9375c0 100644
--- a/plugins/fourier/fourierR.ml
+++ b/plugins/fourier/fourierR.ml
@@ -12,7 +12,7 @@
des inéquations et équations sont entiers. En attendant la tactique Field.
*)
-open Term
+open Constr
open Tactics
open Names
open Globnames
@@ -27,11 +27,7 @@ qui donne le coefficient d'un terme du calcul des constructions,
qui est zéro si le terme n'y est pas.
*)
-module Constrhash = Hashtbl.Make
- (struct type t = constr
- let equal = eq_constr
- let hash = hash_constr
- end)
+module Constrhash = Hashtbl.Make(Constr)
type flin = {fhom: rational Constrhash.t;
fcste:rational};;
@@ -84,7 +80,7 @@ let string_of_R_constant kn =
| _ -> "constant_not_of_R"
let rec string_of_R_constr c =
- match kind_of_term c with
+ match Constr.kind c with
Cast (c,_,_) -> string_of_R_constr c
|Const (c,_) -> string_of_R_constant c
| _ -> "not_of_constant"
@@ -92,7 +88,7 @@ let rec string_of_R_constr c =
exception NoRational
let rec rational_of_constr c =
- match kind_of_term c with
+ match Constr.kind c with
| Cast (c,_,_) -> (rational_of_constr c)
| App (c,args) ->
(match (string_of_R_constr c) with
@@ -125,7 +121,7 @@ exception NoLinear
let rec flin_of_constr c =
try(
- match kind_of_term c with
+ match Constr.kind c with
| Cast (c,_,_) -> (flin_of_constr c)
| App (c,args) ->
(match (string_of_R_constr c) with
@@ -192,9 +188,9 @@ exception NoIneq
let ineq1_of_constr (h,t) =
let h = EConstr.Unsafe.to_constr h in
let t = EConstr.Unsafe.to_constr t in
- match (kind_of_term t) with
+ match (Constr.kind t) with
| App (f,args) ->
- (match kind_of_term f with
+ (match Constr.kind f with
| Const (c,_) when Array.length args = 2 ->
let t1= args.(0) in
let t2= args.(1) in
@@ -233,7 +229,7 @@ let ineq1_of_constr (h,t) =
let t0= args.(0) in
let t1= args.(1) in
let t2= args.(2) in
- (match (kind_of_term t0) with
+ (match (Constr.kind t0) with
| Const (c,_) ->
(match (string_of_R_constant c) with
| "R"->
@@ -438,7 +434,7 @@ let tac_use h =
(*
let is_ineq (h,t) =
- match (kind_of_term t) with
+ match (Constr.kind t) with
App (f,args) ->
(match (string_of_R_constr f) with
"Rlt" -> true
@@ -479,7 +475,7 @@ let rec fourier () =
(* si le but est une inéquation, on introduit son contraire,
et le but à prouver devient False *)
try
- match (kind_of_term goal) with
+ match (Constr.kind goal) with
App (f,args) ->
let get = eget in
(match (string_of_R_constr f) with
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 018b51517..722dbc16b 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -2,6 +2,7 @@ open Printer
open CErrors
open Util
open Term
+open Constr
open Vars
open Namegen
open Names
@@ -80,7 +81,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let is_pte =
let set = List.fold_right Id.Set.add ptes_vars Id.Set.empty in
fun t ->
- match kind_of_term t with
+ match Constr.kind t with
| Var id -> Id.Set.mem id set
| _ -> false
in
@@ -100,13 +101,13 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let pre_princ = EConstr.Unsafe.to_constr pre_princ in
let pre_princ = substl (List.map mkVar ptes_vars) pre_princ in
let is_dom c =
- match kind_of_term c with
+ match Constr.kind c with
| Ind((u,_),_) -> MutInd.equal u rel_as_kn
| Construct(((u,_),_),_) -> MutInd.equal u rel_as_kn
| _ -> false
in
let get_fun_num c =
- match kind_of_term c with
+ match Constr.kind c with
| Ind((_,num),_) -> num
| Construct(((_,num),_),_) -> num
| _ -> assert false
@@ -119,7 +120,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
let rec compute_new_princ_type remove env pre_princ : types*(constr list) =
let (new_princ_type,_) as res =
- match kind_of_term pre_princ with
+ match Constr.kind pre_princ with
| Rel n ->
begin
try match Environ.lookup_rel n env with
@@ -149,12 +150,12 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
let new_f,binders_to_remove_from_f = compute_new_princ_type remove env f in
applistc new_f new_args,
- list_union_eq eq_constr binders_to_remove_from_f binders_to_remove
+ list_union_eq Constr.equal binders_to_remove_from_f binders_to_remove
| LetIn(x,v,t,b) ->
compute_new_princ_type_for_letin remove env x v t b
| _ -> pre_princ,[]
in
-(* let _ = match kind_of_term pre_princ with *)
+(* let _ = match Constr.kind pre_princ with *)
(* | Prod _ -> *)
(* observe(str "compute_new_princ_type for "++ *)
(* pr_lconstr_env env pre_princ ++ *)
@@ -170,13 +171,13 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
let new_env = Environ.push_rel (LocalAssum (x,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
- if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
- then (pop new_b), filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b
+ if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b
+ then (pop new_b), filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b
else
(
bind_fun(new_x,new_t,new_b),
list_union_eq
- eq_constr
+ Constr.equal
binders_to_remove_from_t
(List.map pop binders_to_remove_from_b)
)
@@ -189,7 +190,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Toberemoved_with_rel (n,c) ->
(* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *)
let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in
- new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b)
+ new_b, list_add_set_eq Constr.equal (mkRel n) (List.map pop binders_to_remove_from_b)
end
and compute_new_princ_type_for_letin remove env x v t b =
begin
@@ -199,14 +200,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let new_x : Name.t = get_name (Termops.ids_of_context env) x in
let new_env = Environ.push_rel (LocalDef (x,v,t)) env in
let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in
- if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b
- then (pop new_b),filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b
+ if List.exists (Constr.equal (mkRel 1)) binders_to_remove_from_b
+ then (pop new_b),filter_map (Constr.equal (mkRel 1)) pop binders_to_remove_from_b
else
(
mkLetIn(new_x,new_v,new_t,new_b),
list_union_eq
- eq_constr
- (list_union_eq eq_constr binders_to_remove_from_t binders_to_remove_from_v)
+ Constr.equal
+ (list_union_eq Constr.equal binders_to_remove_from_t binders_to_remove_from_v)
(List.map pop binders_to_remove_from_b)
)
@@ -218,12 +219,12 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Toberemoved_with_rel (n,c) ->
(* observe (str "Decl of "++Ppconstr.Name.print x ++ str " is removed "); *)
let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in
- new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b)
+ new_b, list_add_set_eq Constr.equal (mkRel n) (List.map pop binders_to_remove_from_b)
end
and compute_new_princ_type_with_acc remove env e (c_acc,to_remove_acc) =
let new_e,to_remove_from_e = compute_new_princ_type remove env e
in
- new_e::c_acc,list_union_eq eq_constr to_remove_from_e to_remove_acc
+ new_e::c_acc,list_union_eq Constr.equal to_remove_from_e to_remove_acc
in
(* observe (str "Computing new principe from " ++ pr_lconstr_env env_with_params_and_predicates pre_princ); *)
let pre_res,_ =
@@ -329,7 +330,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
| Some (id) -> id,id
| None ->
let id_of_f = Label.to_id (Constant.label (fst f)) in
- id_of_f,Indrec.make_elimination_ident id_of_f (family_of_sort type_sort)
+ id_of_f,Indrec.make_elimination_ident id_of_f (Sorts.family type_sort)
in
let names = ref [new_princ_name] in
let hook =
@@ -389,7 +390,7 @@ exception Not_Rec
let get_funs_constant mp dp =
let get_funs_constant const e : (Names.Constant.t*int) array =
- match kind_of_term ((strip_lam e)) with
+ match Constr.kind ((strip_lam e)) with
| Fix((_,(na,_,_))) ->
Array.mapi
(fun i na ->
@@ -430,7 +431,7 @@ let get_funs_constant mp dp =
let first_params = List.hd l_params in
List.iter
(fun params ->
- if not (List.equal (fun (n1, c1) (n2, c2) -> Name.equal n1 n2 && eq_constr c1 c2) first_params params)
+ if not (List.equal (fun (n1, c1) (n2, c2) -> Name.equal n1 n2 && Constr.equal c1 c2) first_params params)
then user_err Pp.(str "Not a mutal recursive block")
)
l_params
@@ -439,7 +440,7 @@ let get_funs_constant mp dp =
let _check_bodies =
try
let extract_info is_first body =
- match kind_of_term body with
+ match Constr.kind body with
| Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca)
| _ ->
if is_first && Int.equal (List.length l_bodies) 1
@@ -450,7 +451,7 @@ let get_funs_constant mp dp =
let check body = (* Hope this is correct *)
let eq_infos (ia1, na1, ta1, ca1) (ia2, na2, ta2, ca2) =
Array.equal Int.equal ia1 ia2 && Array.equal Name.equal na1 na2 &&
- Array.equal eq_constr ta1 ta2 && Array.equal eq_constr ca1 ca2
+ Array.equal Constr.equal ta1 ta2 && Array.equal Constr.equal ca1 ca2
in
if not (eq_infos first_infos (extract_info false body))
then user_err Pp.(str "Not a mutal recursive block")
@@ -574,7 +575,7 @@ let make_scheme evd (fas : (pconstant*Sorts.family) list) : Safe_typing.private_
let t = (strip_prod_assum t) in
let applied_g = List.hd (List.rev (snd (decompose_app t))) in
let g = fst (decompose_app applied_g) in
- if eq_constr f g
+ if Constr.equal f g
then raise (Found_type j);
observe (Printer.pr_lconstr f ++ str " <> " ++
Printer.pr_lconstr g)
diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli
index 2eb1b7935..a3315f22c 100644
--- a/plugins/funind/functional_principles_types.mli
+++ b/plugins/funind/functional_principles_types.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
val generate_functional_principle :
Evd.evar_map ref ->
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index e8e5bfccc..8ab6dbcdf 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -2,6 +2,7 @@ open Printer
open Pp
open Names
open Term
+open Constr
open Vars
open Glob_term
open Glob_ops
@@ -1015,7 +1016,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in
observe (str " computing new type for jmeq : done") ;
let new_args =
- match kind_of_term eq'_as_constr with
+ match Constr.kind eq'_as_constr with
| App(_,[|_;_;ty;_|]) ->
let ty = Array.to_list (snd (destApp ty)) in
let ty' = snd (Util.List.chop nparam ty) in
@@ -1297,7 +1298,7 @@ let rec rebuild_return_type rt =
CAst.make @@ Constrexpr.CSort(GType []))
let do_build_inductive
- evd (funconstants: Term.pconstant list) (funsargs: (Name.t * glob_constr * glob_constr option) list list)
+ evd (funconstants: pconstant list) (funsargs: (Name.t * glob_constr * glob_constr option) list list)
returned_types
(rtl:glob_constr list) =
let _time1 = System.get_time () in
diff --git a/plugins/funind/glob_term_to_relation.mli b/plugins/funind/glob_term_to_relation.mli
index 0cab5a6d3..ff0e98d00 100644
--- a/plugins/funind/glob_term_to_relation.mli
+++ b/plugins/funind/glob_term_to_relation.mli
@@ -11,7 +11,7 @@ val build_inductive :
Id.t list -> (* The list of function name *)
*)
Evd.evar_map ->
- Term.pconstant list ->
+ Constr.pconstant list ->
(Name.t*Glob_term.glob_constr*Glob_term.glob_constr option) list list -> (* The list of function args *)
Constrexpr.constr_expr list -> (* The list of function returned type *)
Glob_term.glob_constr list -> (* the list of body *)
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 76fcd5ec8..e9102e9c8 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -1,8 +1,10 @@
open Names
open Pp
+open Constr
open Libnames
open Globnames
open Refiner
+
let mk_prefix pre id = Id.of_string (pre^(Id.to_string id))
let mk_rel_id = mk_prefix "R_"
let mk_correct_id id = Nameops.add_suffix (mk_rel_id id) "_correct"
@@ -111,7 +113,7 @@ let const_of_id id =
(str "cannot find " ++ Id.print id)
let def_of_const t =
- match (Term.kind_of_term t) with
+ match Constr.kind t with
Term.Const sp ->
(try (match Environ.constant_opt_value_in (Global.env()) sp with
| Some c -> c
@@ -330,8 +332,6 @@ let discharge_Function (_,finfos) =
is_general = finfos.is_general
}
-open Term
-
let pr_ocst c =
Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) c (mt ())
@@ -545,9 +545,9 @@ let prodn n env b =
(* compose_prod [xn:Tn;..;x1:T1] b = (x1:T1)..(xn:Tn)b *)
let compose_prod l b = prodn (List.length l) l b
-type tcc_lemma_value =
+type tcc_lemma_value =
| Undefined
- | Value of Term.constr
+ | Value of constr
| Not_needed
(* We only "purify" on exceptions *)
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index d41abee87..5cc7163aa 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -38,7 +38,7 @@ val chop_rlambda_n : int -> Glob_term.glob_constr ->
val chop_rprod_n : int -> Glob_term.glob_constr ->
(Name.t*Glob_term.glob_constr) list * Glob_term.glob_constr
-val def_of_const : Term.constr -> Term.constr
+val def_of_const : Constr.t -> Constr.t
val eq : EConstr.constr Lazy.t
val refl_equal : EConstr.constr Lazy.t
val const_of_id: Id.t -> Globnames.global_reference(* constantyes *)
@@ -118,10 +118,10 @@ val decompose_lam_n : Evd.evar_map -> int -> EConstr.t ->
(Names.Name.t * EConstr.t) list * EConstr.t
val compose_lam : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t
val compose_prod : (Names.Name.t * EConstr.t) list -> EConstr.t -> EConstr.t
-
-type tcc_lemma_value =
+
+type tcc_lemma_value =
| Undefined
- | Value of Term.constr
+ | Value of Constr.t
| Not_needed
val funind_purify : ('a -> 'b) -> ('a -> 'b)
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 93317fce1..692a874d3 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -12,6 +12,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open EConstr
open Vars
open Pp
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 77c26f8ce..b372241d2 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -18,6 +18,7 @@ open Vernacexpr
open Pp
open Names
open Term
+open Constr
open Vars
open Declarations
open Glob_term
@@ -36,19 +37,19 @@ let rec popn i c = if i<=0 then c else pop (popn (i-1) c)
(** Substitutions in constr *)
let compare_constr_nosub t1 t2 =
- if compare_constr (fun _ _ -> false) t1 t2
+ if Constr.compare_head (fun _ _ -> false) t1 t2
then true
else false
let rec compare_constr' t1 t2 =
if compare_constr_nosub t1 t2
then true
- else (compare_constr (compare_constr') t1 t2)
+ else (Constr.compare_head (compare_constr') t1 t2)
let rec substitterm prof t by_t in_u =
if (compare_constr' (lift prof t) in_u)
then (lift prof by_t)
- else map_constr_with_binders succ
+ else Constr.map_with_binders succ
(fun i -> substitterm i t by_t) prof in_u
let lift_ldecl n ldecl = List.map (fun (x,y) -> x,lift n y) ldecl
@@ -954,16 +955,16 @@ let funify_branches relinfo nfuns branch =
| Some (IndRef ((mutual_ind,i) as ind)) -> mutual_ind,ind
| _ -> assert false in
let is_dom c =
- match kind_of_term c with
+ match Constr.kind c with
| Ind(((u,_),_)) | Construct(((u,_),_),_) -> MutInd.equal u mut_induct
| _ -> false in
let _dom_i c =
assert (is_dom c);
- match kind_of_term c with
+ match Constr.kind c with
| Ind((u,i)) | Construct((u,_),i) -> i
| _ -> assert false in
let _is_pred c shift =
- match kind_of_term c with
+ match Constr.kind c with
| Rel i -> let reali = i-shift in (reali>=0 && reali<relinfo.nbranches)
| _ -> false in
(* FIXME: *)
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 76f859ed7..2fdc3bc37 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -10,6 +10,7 @@
module CVars = Vars
open Term
+open Constr
open EConstr
open Vars
open Namegen
@@ -69,7 +70,7 @@ let declare_fun f_id kind ?(ctx=Univ.UContext.empty) value =
let defined () = Lemmas.save_proof (Vernacexpr.(Proved (Transparent,None)))
let def_of_const t =
- match (kind_of_term t) with
+ match (Constr.kind t) with
Const sp ->
(try (match constant_opt_value_in (Global.env ()) sp with
| Some c -> c
@@ -143,7 +144,7 @@ let nat = function () -> (coq_init_constant "nat")
let iter_ref () =
try find_reference ["Recdef"] "iter"
with Not_found -> user_err Pp.(str "module Recdef not loaded")
-let iter = function () -> (constr_of_global (delayed_force iter_ref))
+let iter_rd = function () -> (constr_of_global (delayed_force iter_ref))
let eq = function () -> (coq_init_constant "eq")
let le_lt_SS = function () -> (constant ["Recdef"] "le_lt_SS")
let le_lt_n_Sm = function () -> (coq_constant arith_Lt "le_lt_n_Sm")
@@ -175,8 +176,9 @@ let simpl_iter clause =
clause
(* Others ugly things ... *)
-let (value_f:Term.constr list -> global_reference -> Term.constr) =
+let (value_f: Constr.t list -> global_reference -> Constr.t) =
let open Term in
+ let open Constr in
fun al fterm ->
let rev_x_id_l =
(
@@ -207,7 +209,7 @@ let (value_f:Term.constr list -> global_reference -> Term.constr) =
let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in
it_mkLambda_or_LetIn body context
-let (declare_f : Id.t -> logical_kind -> Term.constr list -> global_reference -> global_reference) =
+let (declare_f : Id.t -> logical_kind -> Constr.t list -> global_reference -> global_reference) =
fun f_id kind input_type fterm_ref ->
declare_fun f_id kind (value_f input_type fterm_ref);;
@@ -1039,11 +1041,12 @@ let prove_eq = travel equation_info
*)
let compute_terminate_type nb_args func =
let open Term in
+ let open Constr in
let open CVars in
let _,a_arrow_b,_ = destLambda(def_of_const (constr_of_global func)) in
let rev_args,b = decompose_prod_n nb_args a_arrow_b in
let left =
- mkApp(delayed_force iter,
+ mkApp(delayed_force iter_rd,
Array.of_list
(lift 5 a_arrow_b:: mkRel 3::
constr_of_global func::mkRel 1::
@@ -1460,7 +1463,7 @@ let start_equation (f:global_reference) (term_f:global_reference)
let (com_eqn : int -> Id.t ->
global_reference -> global_reference -> global_reference
- -> Term.constr -> unit) =
+ -> Constr.t -> unit) =
fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type ->
let open CVars in
let opacity =
@@ -1514,6 +1517,7 @@ let (com_eqn : int -> Id.t ->
let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num eq
generate_induction_principle using_lemmas : unit =
let open Term in
+ let open Constr in
let open CVars in
let env = Global.env() in
let evd = ref (Evd.from_env env) in
@@ -1536,7 +1540,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
(* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *)
(* Pp.msgnl (str "rec_arg_num := " ++ str (string_of_int rec_arg_num)); *)
(* Pp.msgnl (str "eq' := " ++ str (string_of_int rec_arg_num)); *)
- match kind_of_term eq' with
+ match Constr.kind eq' with
| App(e,[|_;_;eq_fix|]) ->
mkLambda (Name function_name,function_type,subst_var function_name (compose_lam res_vars eq_fix))
| _ -> failwith "Recursive Definition (res not eq)"
diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli
index 63bbdbe7e..50b84731b 100644
--- a/plugins/funind/recdef.mli
+++ b/plugins/funind/recdef.mli
@@ -1,3 +1,4 @@
+open Constr
(* val evaluable_of_global_reference : Libnames.global_reference -> Names.evaluable_global_reference *)
val tclUSER_if_not_mes :
@@ -11,9 +12,9 @@ bool ->
Constrintern.internalization_env ->
Constrexpr.constr_expr ->
Constrexpr.constr_expr ->
- int -> Constrexpr.constr_expr -> (Term.pconstant ->
+ int -> Constrexpr.constr_expr -> (pconstant ->
Indfun_common.tcc_lemma_value ref ->
- Term.pconstant ->
- Term.pconstant -> int -> EConstr.types -> int -> EConstr.constr -> 'a) -> Constrexpr.constr_expr list -> unit
+ pconstant ->
+ pconstant -> int -> EConstr.types -> int -> EConstr.constr -> 'a) -> Constrexpr.constr_expr list -> unit
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 65c186a41..4b1555e55 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -415,7 +415,7 @@ VERNAC COMMAND EXTEND DeriveInversionClear
-> [ add_inversion_lemma_exn na c s false inv_clear_tac ]
| [ "Derive" "Inversion_clear" ident(na) "with" constr(c) ] => [ seff na ]
- -> [ add_inversion_lemma_exn na c InProp false inv_clear_tac ]
+ -> [ add_inversion_lemma_exn na c Sorts.InProp false inv_clear_tac ]
END
VERNAC COMMAND EXTEND DeriveInversion
@@ -424,7 +424,7 @@ VERNAC COMMAND EXTEND DeriveInversion
-> [ add_inversion_lemma_exn na c s false inv_tac ]
| [ "Derive" "Inversion" ident(na) "with" constr(c) ] => [ seff na ]
- -> [ add_inversion_lemma_exn na c InProp false inv_tac ]
+ -> [ add_inversion_lemma_exn na c Sorts.InProp false inv_tac ]
END
VERNAC COMMAND EXTEND DeriveDependentInversion
@@ -514,7 +514,7 @@ let cache_transitivity_lemma (_,(left,lem)) =
let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref)
-let inTransitivity : bool * Term.constr -> obj =
+let inTransitivity : bool * Constr.t -> obj =
declare_object {(default_object "TRANSITIVITY-STEPS") with
cache_function = cache_transitivity_lemma;
open_function = (fun i o -> if Int.equal i 1 then cache_transitivity_lemma o);
diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4
index 104977aef..ed2d9da63 100644
--- a/plugins/ltac/g_class.ml4
+++ b/plugins/ltac/g_class.ml4
@@ -91,7 +91,7 @@ END
(** TODO: DEPRECATE *)
(* A progress test that allows to see if the evars have changed *)
-open Term
+open Constr
open Proofview.Notations
let rec eq_constr_mod_evars sigma x y =
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index e467d3e2c..6522fc28f 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -1120,10 +1120,10 @@ let pr_goal_selector ~toplevel s =
let ty = EConstr.Unsafe.to_constr ty in
let rec strip_ty acc n ty =
if n=0 then (List.rev acc, EConstr.of_constr ty) else
- match Term.kind_of_term ty with
- Term.Prod(na,a,b) ->
- strip_ty (([Loc.tag na],EConstr.of_constr a)::acc) (n-1) b
- | _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in
+ match Constr.kind ty with
+ | Constr.Prod(na,a,b) ->
+ strip_ty (([Loc.tag na],EConstr.of_constr a)::acc) (n-1) b
+ | _ -> user_err Pp.(str "Cannot translate fix tactic: not enough products") in
strip_ty [] n ty
let pr_atomic_tactic_level env sigma n t =
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 1809f0fcd..705a51edd 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -12,7 +12,7 @@ open CErrors
open Util
open Nameops
open Namegen
-open Term
+open Constr
open EConstr
open Vars
open Reduction
@@ -426,7 +426,7 @@ let split_head = function
| [] -> assert(false)
let eq_pb (ty, env, x, y as pb) (ty', env', x', y' as pb') =
- pb == pb' || (ty == ty' && Term.eq_constr x x' && Term.eq_constr y y')
+ pb == pb' || (ty == ty' && Constr.equal x x' && Constr.equal y y')
let problem_inclusion x y =
List.for_all (fun pb -> List.exists (fun pb' -> eq_pb pb pb') y) x
@@ -928,8 +928,8 @@ let fold_match ?(force=false) env sigma c =
it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx)
in
let sk =
- if sortp == InProp then
- if sortc == InProp then
+ if sortp == Sorts.InProp then
+ if sortc == Sorts.InProp then
if dep then case_dep_scheme_kind_from_prop
else case_scheme_kind_from_prop
else (
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index 63e891b45..1306c590b 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -37,7 +37,7 @@ type ('constr,'redexpr) strategy_ast =
type rewrite_proof =
| RewPrf of constr * constr
- | RewCast of Term.cast_kind
+ | RewCast of Constr.cast_kind
type evars = evar_map * Evar.Set.t (* goal evars, constraint evars *)
diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml
index 4d171ecbc..c03a86732 100644
--- a/plugins/ltac/taccoerce.ml
+++ b/plugins/ltac/taccoerce.ml
@@ -8,7 +8,7 @@
open Util
open Names
-open Term
+open Constr
open EConstr
open Misctypes
open Genarg
@@ -172,8 +172,8 @@ let id_of_name = function
| Sort s ->
begin
match ESorts.kind sigma s with
- | Prop _ -> Label.to_id (Label.make "Prop")
- | Type _ -> Label.to_id (Label.make "Type")
+ | Sorts.Prop _ -> Label.to_id (Label.make "Prop")
+ | Sorts.Type _ -> Label.to_id (Label.make "Type")
end
| _ -> fail()
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index fc6781b06..218342efe 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -20,6 +20,7 @@ open Pp
open Mutils
open Goptions
open Names
+open Constr
(**
* Debug flag
@@ -580,9 +581,9 @@ struct
| Ukn
| BadStr of string
| BadNum of int
- | BadTerm of Term.constr
+ | BadTerm of constr
| Msg of string
- | Goal of (Term.constr list ) * Term.constr * parse_error
+ | Goal of (constr list ) * constr * parse_error
let string_of_error = function
| Ukn -> "ukn"
@@ -1521,7 +1522,7 @@ let rec witness prover l1 l2 =
let rec apply_ids t ids =
match ids with
| [] -> t
- | i::ids -> apply_ids (Term.mkApp(t,[| Term.mkVar i |])) ids
+ | i::ids -> apply_ids (mkApp(t,[| mkVar i |])) ids
let coq_Node =
lazy (gen_constant_in_modules "VarMap"
diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml
index 72934a15d..559dfab52 100644
--- a/plugins/nsatz/nsatz.ml
+++ b/plugins/nsatz/nsatz.ml
@@ -8,7 +8,7 @@
open CErrors
open Util
-open Term
+open Constr
open Tactics
open Coqlib
@@ -204,42 +204,42 @@ else
mkt_app ttpow [Lazy.force tz; mkt_term t1; mkt_n (num_of_int n)]
let rec parse_pos p =
- match kind_of_term p with
+ match Constr.kind p with
| App (a,[|p2|]) ->
- if eq_constr a (Lazy.force pxO) then num_2 */ (parse_pos p2)
+ if Constr.equal a (Lazy.force pxO) then num_2 */ (parse_pos p2)
else num_1 +/ (num_2 */ (parse_pos p2))
| _ -> num_1
let parse_z z =
- match kind_of_term z with
+ match Constr.kind z with
| App (a,[|p2|]) ->
- if eq_constr a (Lazy.force zpos) then parse_pos p2 else (num_0 -/ (parse_pos p2))
+ if Constr.equal a (Lazy.force zpos) then parse_pos p2 else (num_0 -/ (parse_pos p2))
| _ -> num_0
let parse_n z =
- match kind_of_term z with
+ match Constr.kind z with
| App (a,[|p2|]) ->
parse_pos p2
| _ -> num_0
let rec parse_term p =
- match kind_of_term p with
+ match Constr.kind p with
| App (a,[|_;p2|]) ->
- if eq_constr a (Lazy.force ttvar) then Var (string_of_num (parse_pos p2))
- else if eq_constr a (Lazy.force ttconst) then Const (parse_z p2)
- else if eq_constr a (Lazy.force ttopp) then Opp (parse_term p2)
+ if Constr.equal a (Lazy.force ttvar) then Var (string_of_num (parse_pos p2))
+ else if Constr.equal a (Lazy.force ttconst) then Const (parse_z p2)
+ else if Constr.equal a (Lazy.force ttopp) then Opp (parse_term p2)
else Zero
| App (a,[|_;p2;p3|]) ->
- if eq_constr a (Lazy.force ttadd) then Add (parse_term p2, parse_term p3)
- else if eq_constr a (Lazy.force ttsub) then Sub (parse_term p2, parse_term p3)
- else if eq_constr a (Lazy.force ttmul) then Mul (parse_term p2, parse_term p3)
- else if eq_constr a (Lazy.force ttpow) then
+ if Constr.equal a (Lazy.force ttadd) then Add (parse_term p2, parse_term p3)
+ else if Constr.equal a (Lazy.force ttsub) then Sub (parse_term p2, parse_term p3)
+ else if Constr.equal a (Lazy.force ttmul) then Mul (parse_term p2, parse_term p3)
+ else if Constr.equal a (Lazy.force ttpow) then
Pow (parse_term p2, int_of_num (parse_n p3))
else Zero
| _ -> Zero
let rec parse_request lp =
- match kind_of_term lp with
+ match Constr.kind lp with
| App (_,[|_|]) -> []
| App (_,[|_;p;lp1|]) ->
(parse_term p)::(parse_request lp1)
diff --git a/plugins/nsatz/nsatz.mli b/plugins/nsatz/nsatz.mli
index d6e3071aa..e50a12a50 100644
--- a/plugins/nsatz/nsatz.mli
+++ b/plugins/nsatz/nsatz.mli
@@ -6,4 +6,4 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-val nsatz_compute : Term.constr -> unit Proofview.tactic
+val nsatz_compute : Constr.t -> unit Proofview.tactic
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index e1e73b1c3..96bf31b11 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -166,11 +166,7 @@ exchange ?1 and ?2 in the example above)
*)
-module ConstrSet = Set.Make(
- struct
- type t = Term.constr
- let compare = Term.compare
- end)
+module ConstrSet = Set.Make(Constr)
type inversion_scheme = {
normal_lhs_rhs : (constr * constr_pattern) list;
@@ -385,11 +381,7 @@ let rec sort_subterm gl l =
| [] -> []
| h::t -> insert h (sort_subterm gl t)
-module Constrhash = Hashtbl.Make
- (struct type t = Term.constr
- let equal = Term.eq_constr
- let hash = Term.hash_constr
- end)
+module Constrhash = Hashtbl.Make(Constr)
let subst_meta subst c =
let subst = List.map (fun (i, c) -> i, EConstr.Unsafe.to_constr c) subst in
diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml
index c27ac2ea4..5397b0065 100644
--- a/plugins/romega/const_omega.ml
+++ b/plugins/romega/const_omega.ml
@@ -7,14 +7,16 @@
*************************************************************************)
open Names
+open Term
+open Constr
let module_refl_name = "ReflOmegaCore"
let module_refl_path = ["Coq"; "romega"; module_refl_name]
type result =
| Kvar of string
- | Kapp of string * Term.constr list
- | Kimp of Term.constr * Term.constr
+ | Kapp of string * constr list
+ | Kimp of constr * constr
| Kufo
let meaningful_submodule = [ "Z"; "N"; "Pos" ]
@@ -30,27 +32,27 @@ let string_of_global r =
prefix^(Names.Id.to_string (Nametab.basename_of_global r))
let destructurate t =
- let c, args = Term.decompose_app t in
- match Term.kind_of_term c, args with
- | Term.Const (sp,_), args ->
+ let c, args = decompose_app t in
+ match Constr.kind c, args with
+ | Const (sp,_), args ->
Kapp (string_of_global (Globnames.ConstRef sp), args)
- | Term.Construct (csp,_) , args ->
+ | Construct (csp,_) , args ->
Kapp (string_of_global (Globnames.ConstructRef csp), args)
- | Term.Ind (isp,_), args ->
+ | Ind (isp,_), args ->
Kapp (string_of_global (Globnames.IndRef isp), args)
- | Term.Var id, [] -> Kvar(Names.Id.to_string id)
- | Term.Prod (Anonymous,typ,body), [] -> Kimp(typ,body)
+ | Var id, [] -> Kvar(Names.Id.to_string id)
+ | Prod (Anonymous,typ,body), [] -> Kimp(typ,body)
| _ -> Kufo
exception DestConstApp
let dest_const_apply t =
- let f,args = Term.decompose_app t in
+ let f,args = decompose_app t in
let ref =
- match Term.kind_of_term f with
- | Term.Const (sp,_) -> Globnames.ConstRef sp
- | Term.Construct (csp,_) -> Globnames.ConstructRef csp
- | Term.Ind (isp,_) -> Globnames.IndRef isp
+ match Constr.kind f with
+ | Const (sp,_) -> Globnames.ConstRef sp
+ | Construct (csp,_) -> Globnames.ConstructRef csp
+ | Ind (isp,_) -> Globnames.IndRef isp
| _ -> raise DestConstApp
in Nametab.basename_of_global ref, args
@@ -129,7 +131,7 @@ let coq_O = lazy(init_constant "O")
let rec mk_nat = function
| 0 -> Lazy.force coq_O
- | n -> Term.mkApp (Lazy.force coq_S, [| mk_nat (n-1) |])
+ | n -> mkApp (Lazy.force coq_S, [| mk_nat (n-1) |])
(* Lists *)
@@ -141,47 +143,47 @@ let mkListConst c =
if Global.is_polymorphic r then fun u -> Univ.Instance.of_array [|u|]
else fun _ -> Univ.Instance.empty
in
- fun u -> Term.mkConstructU (Globnames.destConstructRef r, inst u)
+ fun u -> mkConstructU (Globnames.destConstructRef r, inst u)
-let coq_cons univ typ = Term.mkApp (mkListConst "cons" univ, [|typ|])
-let coq_nil univ typ = Term.mkApp (mkListConst "nil" univ, [|typ|])
+let coq_cons univ typ = mkApp (mkListConst "cons" univ, [|typ|])
+let coq_nil univ typ = mkApp (mkListConst "nil" univ, [|typ|])
let mk_list univ typ l =
let rec loop = function
| [] -> coq_nil univ typ
| (step :: l) ->
- Term.mkApp (coq_cons univ typ, [| step; loop l |]) in
+ mkApp (coq_cons univ typ, [| step; loop l |]) in
loop l
let mk_plist =
let type1lev = Universes.new_univ_level (Global.current_dirpath ()) in
- fun l -> mk_list type1lev Term.mkProp l
+ fun l -> mk_list type1lev mkProp l
let mk_list = mk_list Univ.Level.set
type parse_term =
- | Tplus of Term.constr * Term.constr
- | Tmult of Term.constr * Term.constr
- | Tminus of Term.constr * Term.constr
- | Topp of Term.constr
- | Tsucc of Term.constr
+ | Tplus of constr * constr
+ | Tmult of constr * constr
+ | Tminus of constr * constr
+ | Topp of constr
+ | Tsucc of constr
| Tnum of Bigint.bigint
| Tother
type parse_rel =
- | Req of Term.constr * Term.constr
- | Rne of Term.constr * Term.constr
- | Rlt of Term.constr * Term.constr
- | Rle of Term.constr * Term.constr
- | Rgt of Term.constr * Term.constr
- | Rge of Term.constr * Term.constr
+ | Req of constr * constr
+ | Rne of constr * constr
+ | Rlt of constr * constr
+ | Rle of constr * constr
+ | Rgt of constr * constr
+ | Rge of constr * constr
| Rtrue
| Rfalse
- | Rnot of Term.constr
- | Ror of Term.constr * Term.constr
- | Rand of Term.constr * Term.constr
- | Rimp of Term.constr * Term.constr
- | Riff of Term.constr * Term.constr
+ | Rnot of constr
+ | Ror of constr * constr
+ | Rand of constr * constr
+ | Rimp of constr * constr
+ | Riff of constr * constr
| Rother
let parse_logic_rel c = match destructurate c with
@@ -209,29 +211,29 @@ let rec mk_positive n =
if Bigint.equal n Bigint.one then Lazy.force coq_xH
else
let (q,r) = Bigint.euclid n Bigint.two in
- Term.mkApp
+ mkApp
((if Bigint.equal r Bigint.zero
then Lazy.force coq_xO else Lazy.force coq_xI),
[| mk_positive q |])
let mk_N = function
| 0 -> Lazy.force coq_N0
- | n -> Term.mkApp (Lazy.force coq_Npos,
+ | n -> mkApp (Lazy.force coq_Npos,
[| mk_positive (Bigint.of_int n) |])
module type Int = sig
- val typ : Term.constr Lazy.t
- val is_int_typ : [ `NF ] Proofview.Goal.t -> Term.constr -> bool
- val plus : Term.constr Lazy.t
- val mult : Term.constr Lazy.t
- val opp : Term.constr Lazy.t
- val minus : Term.constr Lazy.t
-
- val mk : Bigint.bigint -> Term.constr
- val parse_term : Term.constr -> parse_term
- val parse_rel : [ `NF ] Proofview.Goal.t -> Term.constr -> parse_rel
+ val typ : constr Lazy.t
+ val is_int_typ : [ `NF ] Proofview.Goal.t -> constr -> bool
+ val plus : constr Lazy.t
+ val mult : constr Lazy.t
+ val opp : constr Lazy.t
+ val minus : constr Lazy.t
+
+ val mk : Bigint.bigint -> constr
+ val parse_term : constr -> parse_term
+ val parse_rel : [ `NF ] Proofview.Goal.t -> constr -> parse_rel
(* check whether t is built only with numbers and + * - *)
- val get_scalar : Term.constr -> Bigint.bigint option
+ val get_scalar : constr -> Bigint.bigint option
end
module Z : Int = struct
@@ -266,9 +268,9 @@ let recognize_Z t =
let mk_Z n =
if Bigint.equal n Bigint.zero then Lazy.force coq_Z0
else if Bigint.is_strictly_pos n then
- Term.mkApp (Lazy.force coq_Zpos, [| mk_positive n |])
+ mkApp (Lazy.force coq_Zpos, [| mk_positive n |])
else
- Term.mkApp (Lazy.force coq_Zneg, [| mk_positive (Bigint.neg n) |])
+ mkApp (Lazy.force coq_Zneg, [| mk_positive (Bigint.neg n) |])
let mk = mk_Z
diff --git a/plugins/romega/const_omega.mli b/plugins/romega/const_omega.mli
index 80e00e4e1..5ba063d9d 100644
--- a/plugins/romega/const_omega.mli
+++ b/plugins/romega/const_omega.mli
@@ -8,116 +8,117 @@
(** Coq objects used in romega *)
+open Constr
(* from Logic *)
-val coq_refl_equal : Term.constr lazy_t
-val coq_and : Term.constr lazy_t
-val coq_not : Term.constr lazy_t
-val coq_or : Term.constr lazy_t
-val coq_True : Term.constr lazy_t
-val coq_False : Term.constr lazy_t
-val coq_I : Term.constr lazy_t
+val coq_refl_equal : constr lazy_t
+val coq_and : constr lazy_t
+val coq_not : constr lazy_t
+val coq_or : constr lazy_t
+val coq_True : constr lazy_t
+val coq_False : constr lazy_t
+val coq_I : constr lazy_t
(* from ReflOmegaCore/ZOmega *)
-val coq_t_int : Term.constr lazy_t
-val coq_t_plus : Term.constr lazy_t
-val coq_t_mult : Term.constr lazy_t
-val coq_t_opp : Term.constr lazy_t
-val coq_t_minus : Term.constr lazy_t
-val coq_t_var : Term.constr lazy_t
-
-val coq_proposition : Term.constr lazy_t
-val coq_p_eq : Term.constr lazy_t
-val coq_p_leq : Term.constr lazy_t
-val coq_p_geq : Term.constr lazy_t
-val coq_p_lt : Term.constr lazy_t
-val coq_p_gt : Term.constr lazy_t
-val coq_p_neq : Term.constr lazy_t
-val coq_p_true : Term.constr lazy_t
-val coq_p_false : Term.constr lazy_t
-val coq_p_not : Term.constr lazy_t
-val coq_p_or : Term.constr lazy_t
-val coq_p_and : Term.constr lazy_t
-val coq_p_imp : Term.constr lazy_t
-val coq_p_prop : Term.constr lazy_t
-
-val coq_s_bad_constant : Term.constr lazy_t
-val coq_s_divide : Term.constr lazy_t
-val coq_s_not_exact_divide : Term.constr lazy_t
-val coq_s_sum : Term.constr lazy_t
-val coq_s_merge_eq : Term.constr lazy_t
-val coq_s_split_ineq : Term.constr lazy_t
-
-val coq_direction : Term.constr lazy_t
-val coq_d_left : Term.constr lazy_t
-val coq_d_right : Term.constr lazy_t
-
-val coq_e_split : Term.constr lazy_t
-val coq_e_extract : Term.constr lazy_t
-val coq_e_solve : Term.constr lazy_t
-
-val coq_interp_sequent : Term.constr lazy_t
-val coq_do_omega : Term.constr lazy_t
-
-val mk_nat : int -> Term.constr
-val mk_N : int -> Term.constr
+val coq_t_int : constr lazy_t
+val coq_t_plus : constr lazy_t
+val coq_t_mult : constr lazy_t
+val coq_t_opp : constr lazy_t
+val coq_t_minus : constr lazy_t
+val coq_t_var : constr lazy_t
+
+val coq_proposition : constr lazy_t
+val coq_p_eq : constr lazy_t
+val coq_p_leq : constr lazy_t
+val coq_p_geq : constr lazy_t
+val coq_p_lt : constr lazy_t
+val coq_p_gt : constr lazy_t
+val coq_p_neq : constr lazy_t
+val coq_p_true : constr lazy_t
+val coq_p_false : constr lazy_t
+val coq_p_not : constr lazy_t
+val coq_p_or : constr lazy_t
+val coq_p_and : constr lazy_t
+val coq_p_imp : constr lazy_t
+val coq_p_prop : constr lazy_t
+
+val coq_s_bad_constant : constr lazy_t
+val coq_s_divide : constr lazy_t
+val coq_s_not_exact_divide : constr lazy_t
+val coq_s_sum : constr lazy_t
+val coq_s_merge_eq : constr lazy_t
+val coq_s_split_ineq : constr lazy_t
+
+val coq_direction : constr lazy_t
+val coq_d_left : constr lazy_t
+val coq_d_right : constr lazy_t
+
+val coq_e_split : constr lazy_t
+val coq_e_extract : constr lazy_t
+val coq_e_solve : constr lazy_t
+
+val coq_interp_sequent : constr lazy_t
+val coq_do_omega : constr lazy_t
+
+val mk_nat : int -> constr
+val mk_N : int -> constr
(** Precondition: the type of the list is in Set *)
-val mk_list : Term.constr -> Term.constr list -> Term.constr
-val mk_plist : Term.types list -> Term.types
+val mk_list : constr -> constr list -> constr
+val mk_plist : types list -> types
(** Analyzing a coq term *)
(* The generic result shape of the analysis of a term.
One-level depth, except when a number is found *)
type parse_term =
- Tplus of Term.constr * Term.constr
- | Tmult of Term.constr * Term.constr
- | Tminus of Term.constr * Term.constr
- | Topp of Term.constr
- | Tsucc of Term.constr
+ Tplus of constr * constr
+ | Tmult of constr * constr
+ | Tminus of constr * constr
+ | Topp of constr
+ | Tsucc of constr
| Tnum of Bigint.bigint
| Tother
(* The generic result shape of the analysis of a relation.
One-level depth. *)
type parse_rel =
- Req of Term.constr * Term.constr
- | Rne of Term.constr * Term.constr
- | Rlt of Term.constr * Term.constr
- | Rle of Term.constr * Term.constr
- | Rgt of Term.constr * Term.constr
- | Rge of Term.constr * Term.constr
+ Req of constr * constr
+ | Rne of constr * constr
+ | Rlt of constr * constr
+ | Rle of constr * constr
+ | Rgt of constr * constr
+ | Rge of constr * constr
| Rtrue
| Rfalse
- | Rnot of Term.constr
- | Ror of Term.constr * Term.constr
- | Rand of Term.constr * Term.constr
- | Rimp of Term.constr * Term.constr
- | Riff of Term.constr * Term.constr
+ | Rnot of constr
+ | Ror of constr * constr
+ | Rand of constr * constr
+ | Rimp of constr * constr
+ | Riff of constr * constr
| Rother
(* A module factorizing what we should now about the number representation *)
module type Int =
sig
(* the coq type of the numbers *)
- val typ : Term.constr Lazy.t
+ val typ : constr Lazy.t
(* Is a constr expands to the type of these numbers *)
- val is_int_typ : [ `NF ] Proofview.Goal.t -> Term.constr -> bool
+ val is_int_typ : [ `NF ] Proofview.Goal.t -> constr -> bool
(* the operations on the numbers *)
- val plus : Term.constr Lazy.t
- val mult : Term.constr Lazy.t
- val opp : Term.constr Lazy.t
- val minus : Term.constr Lazy.t
+ val plus : constr Lazy.t
+ val mult : constr Lazy.t
+ val opp : constr Lazy.t
+ val minus : constr Lazy.t
(* building a coq number *)
- val mk : Bigint.bigint -> Term.constr
+ val mk : Bigint.bigint -> constr
(* parsing a term (one level, except if a number is found) *)
- val parse_term : Term.constr -> parse_term
+ val parse_term : constr -> parse_term
(* parsing a relation expression, including = < <= >= > *)
- val parse_rel : [ `NF ] Proofview.Goal.t -> Term.constr -> parse_rel
+ val parse_rel : [ `NF ] Proofview.Goal.t -> constr -> parse_rel
(* Is a particular term only made of numbers and + * - ? *)
- val get_scalar : Term.constr -> Bigint.bigint option
+ val get_scalar : constr -> Bigint.bigint option
end
(* Currently, we only use Z numbers *)
diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml
index 661485aee..430b608f4 100644
--- a/plugins/romega/refl_omega.ml
+++ b/plugins/romega/refl_omega.ml
@@ -8,6 +8,7 @@
open Pp
open Util
+open Constr
open Const_omega
module OmegaSolver = Omega_plugin.Omega.MakeOmegaSolver (Bigint)
open OmegaSolver
@@ -27,8 +28,6 @@ let pp i = print_int i; print_newline (); flush stdout
(* More readable than the prefix notation *)
let (>>) = Tacticals.New.tclTHEN
-let mkApp = Term.mkApp
-
(* \section{Types}
\subsection{How to walk in a term}
To represent how to get to a proposition. Only choice points are
@@ -68,14 +67,14 @@ type comparaison = Eq | Leq | Geq | Gt | Lt | Neq
(it could contains some [Term.Var] but no [Term.Rel]). So no need to
lift when breaking or creating arrows. *)
type oproposition =
- Pequa of Term.constr * oequation (* constr = copy of the Coq formula *)
+ Pequa of constr * oequation (* constr = copy of the Coq formula *)
| Ptrue
| Pfalse
| Pnot of oproposition
| Por of int * oproposition * oproposition
| Pand of int * oproposition * oproposition
| Pimp of int * oproposition * oproposition
- | Pprop of Term.constr
+ | Pprop of constr
(* The equations *)
and oequation = {
@@ -102,9 +101,9 @@ and oequation = {
type environment = {
(* La liste des termes non reifies constituant l'environnement global *)
- mutable terms : Term.constr list;
+ mutable terms : constr list;
(* La meme chose pour les propositions *)
- mutable props : Term.constr list;
+ mutable props : constr list;
(* Traduction des indices utilisés ici en les indices finaux utilisés par
* la tactique Omega après dénombrement des variables utiles *)
real_indices : int IntHtbl.t;
@@ -219,7 +218,7 @@ let display_omega_var i = Printf.sprintf "OV%d" i
calcul des variables utiles. *)
let add_reified_atom t env =
- try List.index0 Term.eq_constr t env.terms
+ try List.index0 Constr.equal t env.terms
with Not_found ->
let i = List.length env.terms in
env.terms <- env.terms @ [t]; i
@@ -237,7 +236,7 @@ let set_reified_atom v t env =
(* \subsection{Gestion de l'environnement de proposition pour Omega} *)
(* ajout d'une proposition *)
let add_prop env t =
- try List.index0 Term.eq_constr t env.props
+ try List.index0 Constr.equal t env.props
with Not_found ->
let i = List.length env.props in env.props <- env.props @ [t]; i
@@ -560,7 +559,7 @@ let reify_hyp env gl i =
| LocalDef (_,d,t) when Z.is_int_typ gl (EConstr.Unsafe.to_constr t) ->
let d = EConstr.Unsafe.to_constr d in
let dummy = Lazy.force coq_True in
- let p = mk_equation env ctxt dummy Eq (Term.mkVar i) d in
+ let p = mk_equation env ctxt dummy Eq (mkVar i) d in
i,Defined,p
| LocalDef (_,_,t) | LocalAssum (_,t) ->
let t = EConstr.Unsafe.to_constr t in
@@ -1012,7 +1011,7 @@ let resolution unsafe env (reified_concl,reified_hyps) systems_list =
(fun id ->
match Id.Map.find id reified_hyps with
| Defined,p ->
- reified_of_proposition env p, mk_refl (Term.mkVar id)
+ reified_of_proposition env p, mk_refl (mkVar id)
| Assumed,p ->
reified_of_proposition env (maximize_prop useful_equa_ids p),
EConstr.mkVar id
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 9f02388c3..150c253a7 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -13,6 +13,7 @@ open Ltac_plugin
open CErrors
open Util
open Term
+open Constr
open Tacmach
open Proof_search
open Context.Named.Declaration
@@ -82,7 +83,7 @@ let make_atom atom_env term=
let term = EConstr.Unsafe.to_constr term in
try
let (_,i)=
- List.find (fun (t,_)-> eq_constr term t) atom_env.env
+ List.find (fun (t,_)-> Constr.equal term t) atom_env.env
in Atom i
with Not_found ->
let i=atom_env.next in
diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli
index bec18f6df..b2285a4a1 100644
--- a/plugins/rtauto/refl_tauto.mli
+++ b/plugins/rtauto/refl_tauto.mli
@@ -10,7 +10,7 @@
type atom_env=
{mutable next:int;
- mutable env:(Term.constr*int) list}
+ mutable env:(Constr.t*int) list}
val make_form : atom_env ->
Goal.goal Evd.sigma -> EConstr.types -> Proof_search.form
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index b8fae2494..9e4b896f8 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -10,7 +10,7 @@ open Ltac_plugin
open Pp
open Util
open Names
-open Term
+open Constr
open EConstr
open Vars
open CClosure
@@ -58,13 +58,13 @@ let rec mk_clos_but f_map subs t =
match f_map (global_of_constr_nofail t) with
| Some map -> tag_arg (mk_clos_but f_map subs) map subs (-1) t
| None ->
- (match kind_of_term t with
+ (match Constr.kind t with
App(f,args) -> mk_clos_app_but f_map subs f args 0
| Prod _ -> mk_clos_deep (mk_clos_but f_map) subs t
| _ -> mk_atom t)
and mk_clos_app_but f_map subs f args n =
- let open Term in
+ let open Constr in
if n >= Array.length args then mk_atom(mkApp(f, args))
else
let fargs, args' = Array.chop n args in
@@ -151,7 +151,7 @@ let ic_unsafe c = (*FIXME remove *)
EConstr.of_constr (fst (Constrintern.interp_constr env sigma c))
let decl_constant na ctx c =
- let open Term in
+ let open Constr in
let vars = Univops.universes_of_constr c in
let ctx = Univops.restrict_universe_context (Univ.ContextSet.of_context ctx) vars in
mkConst(declare_constant (Id.of_string na)
@@ -346,11 +346,7 @@ let _ = add_map "ring"
let pr_constr c = pr_econstr c
-module M = struct
- type t = Term.constr
- let compare = Term.compare
-end
-module Cmap = Map.Make(M)
+module Cmap = Map.Make(Constr)
let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table"
let from_name = Summary.ref Spmap.empty ~name:"ring-tac-name-table"
@@ -395,7 +391,7 @@ let subst_th (subst,th) =
let posttac'= Tacsubst.subst_tactic subst th.ring_post_tac in
if c' == th.ring_carrier &&
eq' == th.ring_req &&
- Term.eq_constr set' th.ring_setoid &&
+ Constr.equal set' th.ring_setoid &&
ext' == th.ring_ext &&
morph' == th.ring_morph &&
th' == th.ring_th &&
@@ -933,7 +929,7 @@ let field_equality evd r inv req =
inv_m_lem
let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv =
- let open Term in
+ let open Constr in
check_required_library (cdir@["Field_tac"]);
let (sigma,fth) = ic fth in
let env = Global.env() in
diff --git a/plugins/setoid_ring/newring_ast.mli b/plugins/setoid_ring/newring_ast.mli
index d37582bd7..c26fcc8d1 100644
--- a/plugins/setoid_ring/newring_ast.mli
+++ b/plugins/setoid_ring/newring_ast.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Libnames
open Constrexpr
open Tacexpr
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 1f2d02093..c1d7e6278 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -12,6 +12,7 @@ open Util
open Names
open Evd
open Term
+open Constr
open Termops
open Printer
open Locusops
@@ -465,7 +466,6 @@ let ssrevaltac ist gtac =
(* but stripping global ones. We use the variable names to encode the *)
(* the number of dependencies, so that the transformation is reversible. *)
-open Term
let env_size env = List.length (Environ.named_context env)
let pf_concl gl = EConstr.Unsafe.to_constr (pf_concl gl)
@@ -491,23 +491,23 @@ let pf_abs_evars2 gl rigid (sigma, c0) =
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
nf_evar sigma t in
- let rec put evlist c = match kind_of_term c with
+ let rec put evlist c = match Constr.kind c with
| Evar (k, a) ->
if List.mem_assoc k evlist || Evd.mem sigma0 k || List.mem k rigid then evlist else
let n = max 0 (Array.length a - nenv) in
let t = abs_evar n k in (k, (n, t)) :: put evlist t
- | _ -> fold_constr put evlist c in
+ | _ -> Constr.fold put evlist c in
let evlist = put [] c0 in
if evlist = [] then 0, EConstr.of_constr c0,[], ucst else
let rec lookup k i = function
| [] -> 0, 0
| (k', (n, _)) :: evl -> if k = k' then i, n else lookup k (i + 1) evl in
- let rec get i c = match kind_of_term c with
+ let rec get i c = match Constr.kind c with
| Evar (ev, a) ->
let j, n = lookup ev i evlist in
- if j = 0 then map_constr (get i) c else if n = 0 then mkRel j else
+ if j = 0 then Constr.map (get i) c else if n = 0 then mkRel j else
mkApp (mkRel j, Array.init n (fun k -> get i a.(n - 1 - k)))
- | _ -> map_constr_with_binders ((+) 1) get i c in
+ | _ -> Constr.map_with_binders ((+) 1) get i c in
let rec loop c i = function
| (_, (n, t)) :: evl ->
loop (mkLambda (mk_evar_name n, get (i - 1) t, c)) (i - 1) evl
@@ -551,7 +551,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
| NamedDecl.LocalAssum (x,t) -> mkNamedProd x t c in
let t = Context.Named.fold_inside abs_dc ~init:evi.evar_concl dc in
nf_evar sigma0 (nf_evar sigma t) in
- let rec put evlist c = match kind_of_term c with
+ let rec put evlist c = match Constr.kind c with
| Evar (k, a) ->
if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else
let n = max 0 (Array.length a - nenv) in
@@ -560,7 +560,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
(pf_env gl) sigma (EConstr.of_constr (Evd.evar_concl (Evd.find sigma k))) in
let is_prop = k_ty = InProp in
let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t
- | _ -> fold_constr put evlist c in
+ | _ -> Constr.fold put evlist c in
let evlist = put [] c0 in
if evlist = [] then 0, c0 else
let pr_constr t = Printer.pr_econstr (Reductionops.nf_beta (project gl) (EConstr.of_constr t)) in
@@ -588,17 +588,17 @@ let pf_abs_evars_pirrel gl (sigma, c0) =
let rec lookup k i = function
| [] -> 0, 0
| (k', (n,_,_)) :: evl -> if k = k' then i,n else lookup k (i + 1) evl in
- let rec get evlist i c = match kind_of_term c with
+ let rec get evlist i c = match Constr.kind c with
| Evar (ev, a) ->
let j, n = lookup ev i evlist in
- if j = 0 then map_constr (get evlist i) c else if n = 0 then mkRel j else
+ if j = 0 then Constr.map (get evlist i) c else if n = 0 then mkRel j else
mkApp (mkRel j, Array.init n (fun k -> get evlist i a.(n - 1 - k)))
- | _ -> map_constr_with_binders ((+) 1) (get evlist) i c in
+ | _ -> Constr.map_with_binders ((+) 1) (get evlist) i c in
let rec app extra_args i c = match decompose_app c with
| hd, args when isRel hd && destRel hd = i ->
let j = destRel hd in
mkApp (mkRel j, Array.of_list (List.map (Vars.lift (i-1)) extra_args @ args))
- | _ -> map_constr_with_binders ((+) 1) (app extra_args) i c in
+ | _ -> Constr.map_with_binders ((+) 1) (app extra_args) i c in
let rec loopP evlist c i = function
| (_, (n, t, _)) :: evl ->
let t = get evlist (i - 1) t in
@@ -645,7 +645,7 @@ let pf_abs_cterm gl n c0 =
let c0 = EConstr.Unsafe.to_constr c0 in
let noargs = [|0|] in
let eva = Array.make n noargs in
- let rec strip i c = match kind_of_term c with
+ let rec strip i c = match Constr.kind c with
| App (f, a) when isRel f ->
let j = i - destRel f in
if j >= n || eva.(j) = noargs then mkApp (f, Array.map (strip i) a) else
@@ -653,8 +653,8 @@ let pf_abs_cterm gl n c0 =
let nd = Array.length dp - 1 in
let mkarg k = strip i a.(if k < nd then dp.(k + 1) - j else k + dp.(0)) in
mkApp (f, Array.init (Array.length a - dp.(0)) mkarg)
- | _ -> map_constr_with_binders ((+) 1) strip i c in
- let rec strip_ndeps j i c = match kind_of_term c with
+ | _ -> Constr.map_with_binders ((+) 1) strip i c in
+ let rec strip_ndeps j i c = match Constr.kind c with
| Prod (x, t, c1) when i < j ->
let dl, c2 = strip_ndeps j (i + 1) c1 in
if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else
@@ -665,7 +665,7 @@ let pf_abs_cterm gl n c0 =
if Vars.noccurn 1 c2 then dl, Vars.lift (-1) c2 else
i :: dl, mkLetIn (x, strip i b, strip i t, c2)
| _ -> [], strip i c in
- let rec strip_evars i c = match kind_of_term c with
+ let rec strip_evars i c = match Constr.kind c with
| Lambda (x, t1, c1) when i < n ->
let na = nb_evar_deps x in
let dl, t2 = strip_ndeps (i + na) i t1 in
@@ -760,7 +760,7 @@ let clear_with_wilds wilds clr0 gl =
let id = NamedDecl.get_id nd in
if List.mem id clr || not (List.mem id wilds) then clr else
let vars = Termops.global_vars_set_of_decl (pf_env gl) (project gl) nd in
- let occurs id' = Idset.mem id' vars in
+ let occurs id' = Id.Set.mem id' vars in
if List.exists occurs clr then id :: clr else clr in
Proofview.V82.of_tactic (Tactics.clear (Context.Named.fold_inside extend_clr ~init:clr0 (Tacmach.pf_hyps gl))) gl
diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli
index 2eadd5f26..c39945194 100644
--- a/plugins/ssr/ssrcommon.mli
+++ b/plugins/ssr/ssrcommon.mli
@@ -190,7 +190,7 @@ val pf_merge_uc_of :
val constr_name : evar_map -> EConstr.t -> Name.t
val pf_type_of :
Goal.goal Evd.sigma ->
- Term.constr -> Goal.goal Evd.sigma * Term.types
+ Constr.constr -> Goal.goal Evd.sigma * Constr.types
val pfe_type_of :
Goal.goal Evd.sigma ->
EConstr.t -> Goal.goal Evd.sigma * EConstr.types
@@ -220,7 +220,7 @@ val new_wild_id : tac_ctx -> Names.Id.t * tac_ctx
val pf_fresh_global :
Globnames.global_reference ->
Goal.goal Evd.sigma ->
- Term.constr * Goal.goal Evd.sigma
+ Constr.constr * Goal.goal Evd.sigma
val is_discharged_id : Id.t -> bool
val mk_discharged_id : Id.t -> Id.t
@@ -232,7 +232,7 @@ val new_tmp_id :
val mk_anon_id : string -> Goal.goal Evd.sigma -> Id.t
val pf_abs_evars_pirrel :
Goal.goal Evd.sigma ->
- evar_map * Term.constr -> int * Term.constr
+ evar_map * Constr.constr -> int * Constr.constr
val pf_nbargs : Goal.goal Evd.sigma -> EConstr.t -> int
val gen_tmp_ids :
?ist:Geninterp.interp_sign ->
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 95ca6f49a..e82f222b9 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -11,13 +11,14 @@
open Ltac_plugin
open Util
open Names
+open Term
+open Constr
open Vars
open Locus
open Printer
open Globnames
open Termops
open Tacinterp
-open Term
open Ssrmatching_plugin
open Ssrmatching
@@ -316,7 +317,7 @@ let rw_progress rhs lhs ise = not (EConstr.eq_constr ise lhs (Evarutil.nf_evar i
(* such a generic Leibnitz equation -- short of inspecting the type *)
(* of the elimination lemmas. *)
-let rec strip_prod_assum c = match Term.kind_of_term c with
+let rec strip_prod_assum c = match Constr.kind c with
| Prod (_, _, c') -> strip_prod_assum c'
| LetIn (_, v, _, c') -> strip_prod_assum (subst1 v c)
| Cast (c', _, _) -> strip_prod_assum c'
diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml
index d01bdc1b9..29e96ec59 100644
--- a/plugins/ssr/ssrfwd.ml
+++ b/plugins/ssr/ssrfwd.ml
@@ -32,6 +32,7 @@ let ssrposetac ist (id, (_, t)) gl =
open Pp
open Term
+open Constr
let ssrsettac ist id ((_, (pat, pty)), (_, occ)) gl =
let pat = interp_cpattern ist gl pat (Option.map snd pty) in
@@ -59,10 +60,10 @@ let rec is_Evar_or_CastedMeta sigma x =
(EConstr.isCast sigma x && is_Evar_or_CastedMeta sigma (pi1 (EConstr.destCast sigma x)))
let occur_existential_or_casted_meta c =
- let rec occrec c = match kind_of_term c with
+ let rec occrec c = match Constr.kind c with
| Evar _ -> raise Not_found
| Cast (m,_,_) when isMeta m -> raise Not_found
- | _ -> iter_constr occrec c
+ | _ -> Constr.iter occrec c
in try occrec c; false with Not_found -> true
open Printer
@@ -84,12 +85,12 @@ let pf_find_abstract_proof check_lock gl abstract_n =
let fire gl t = EConstr.Unsafe.to_constr (Reductionops.nf_evar (project gl) (EConstr.of_constr t)) in
let abstract, gl = pf_mkSsrConst "abstract" gl in
let l = Evd.fold_undefined (fun e ei l ->
- match kind_of_term ei.Evd.evar_concl with
+ match Constr.kind ei.Evd.evar_concl with
| App(hd, [|ty; n; lock|])
when (not check_lock ||
(occur_existential_or_casted_meta (fire gl ty) &&
is_Evar_or_CastedMeta (project gl) (EConstr.of_constr @@ fire gl lock))) &&
- Term.eq_constr hd (EConstr.Unsafe.to_constr abstract) && Term.eq_constr n abstract_n -> e::l
+ Constr.equal hd (EConstr.Unsafe.to_constr abstract) && Constr.equal n abstract_n -> e::l
| _ -> l) (project gl) [] in
match l with
| [e] -> e
@@ -286,7 +287,7 @@ let ssrabstract ist gens (*last*) gl =
let p = mkApp (proj2,[|ty;concl;p|]) in
let concl = mkApp(prod,[|ty; concl|]) in
pf_unify_HO gl concl t, p
- | App(hd, [|left; right|]) when Term.eq_constr hd prod ->
+ | App(hd, [|left; right|]) when Term.Constr.equal hd prod ->
find_hole (mkApp (proj1,[|left;right;p|])) left
*)
| _ -> errorstrm(strbrk"abstract constant "++pr_econstr abstract_n++
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index 507b4631b..36dce37ae 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -9,7 +9,8 @@
(* This file is (C) Copyright 2006-2015 Microsoft Corporation and Inria. *)
open Names
-open Term
+module CoqConstr = Constr
+open CoqConstr
open Termops
open Constrexpr
open Constrexpr_ops
@@ -364,7 +365,7 @@ let coerce_search_pattern_to_sort hpat =
let interp_head_pat hpat =
let filter_head, p = coerce_search_pattern_to_sort hpat in
- let rec loop c = match kind_of_term c with
+ let rec loop c = match CoqConstr.kind c with
| Cast (c', _, _) -> loop c'
| Prod (_, _, c') -> loop c'
| LetIn (_, _, _, c') -> loop c'
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index e3e34616b..d5c9e4988 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -18,10 +18,13 @@ let frozen_lexer = CLexer.get_keyword_state () ;;
open Ltac_plugin
open Names
open Pp
-open Pcoq
open Genarg
open Stdarg
open Term
+module CoqConstr = Constr
+open CoqConstr
+open Pcoq
+open Pcoq.Constr
open Vars
open Libnames
open Tactics
@@ -35,10 +38,8 @@ open Evd
open Tacexpr
open Tacinterp
open Pretyping
-open Constr
open Ppconstr
open Printer
-
open Globnames
open Misctypes
open Decl_kinds
@@ -73,7 +74,7 @@ let pp s = !pp_ref s
(** Utils {{{ *****************************************************************)
let env_size env = List.length (Environ.named_context env)
let safeDestApp c =
- match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |]
+ match kind c with App (f, a) -> f, a | _ -> c, [| |]
(* Toplevel constr must be globalized twice ! *)
let glob_constr ist genv = function
| _, Some ce ->
@@ -325,7 +326,7 @@ let unif_FO env ise p c =
let nf_open_term sigma0 ise c =
let c = EConstr.Unsafe.to_constr c in
let s = ise and s' = ref sigma0 in
- let rec nf c' = match kind_of_term c' with
+ let rec nf c' = match kind c' with
| Evar ex ->
begin try nf (existential_value s ex) with _ ->
let k, a = ex in let a' = Array.map nf a in
@@ -333,7 +334,7 @@ let nf_open_term sigma0 ise c =
s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k));
mkEvar (k, a')
end
- | _ -> map_constr nf c' in
+ | _ -> map nf c' in
let copy_def k evi () =
if evar_body evi != Evd.Evar_empty then () else
match Evd.evar_body (Evd.find s k) with
@@ -365,7 +366,7 @@ let pf_unify_HO gl t1 t2 =
re_sig si sigma
(* This is what the definition of iter_constr should be... *)
-let iter_constr_LR f c = match kind_of_term c with
+let iter_constr_LR f c = match kind c with
| Evar (k, a) -> Array.iter f a
| Cast (cc, _, t) -> f cc; f t
| Prod (_, t, b) | Lambda (_, t, b) -> f t; f b
@@ -418,14 +419,14 @@ let all_ok _ _ = true
let proj_nparams c =
try 1 + Recordops.find_projection_nparams (ConstRef c) with _ -> 0
-let isRigid c = match kind_of_term c with
+let isRigid c = match kind c with
| Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true
| _ -> false
let hole_var = mkVar (Id.of_string "_")
let pr_constr_pat c0 =
let rec wipe_evar c =
- if isEvar c then hole_var else map_constr wipe_evar c in
+ if isEvar c then hole_var else map wipe_evar c in
pr_constr (wipe_evar c0)
(* Turn (new) evars into metas *)
@@ -433,11 +434,11 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
let ise = ref ise0 in
let sigma = ref ise0 in
let nenv = env_size env + if hack then 1 else 0 in
- let rec put c = match kind_of_term c with
+ let rec put c = match kind c with
| Evar (k, a as ex) ->
begin try put (existential_value !sigma ex)
with NotInstantiatedEvar ->
- if Evd.mem sigma0 k then map_constr put c else
+ if Evd.mem sigma0 k then map put c else
let evi = Evd.find !sigma k in
let dc = List.firstn (max 0 (Array.length a - nenv)) (evar_filtered_context evi) in
let abs_dc (d, c) = function
@@ -452,7 +453,7 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
sigma := Evd.define k (applistc (mkMeta m) a) !sigma;
put (existential_value !sigma ex)
end
- | _ -> map_constr put c in
+ | _ -> map put c in
let c1 = put c0 in !ise, c1
(* Compile a match pattern from a term; t is the term to fill. *)
@@ -462,7 +463,7 @@ let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p =
let f, a = Reductionops.whd_betaiota_stack ise (EConstr.of_constr p) in
let f = EConstr.Unsafe.to_constr f in
let a = List.map EConstr.Unsafe.to_constr a in
- match kind_of_term f with
+ match kind f with
| Const (p,_) ->
let np = proj_nparams p in
if np = 0 || np > List.length a then KpatConst, f, a else
@@ -490,7 +491,7 @@ let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p =
(* kind and arity for Proj and Flex patterns. *)
let ungen_upat lhs (sigma, uc, t) u =
let f, a = safeDestApp lhs in
- let k = match kind_of_term f with
+ let k = match kind f with
| Var _ | Ind _ | Construct _ -> KpatFixed
| Const _ -> KpatConst
| Evar (k, _) -> if is_defined sigma k then raise NoMatch else KpatEvar k
@@ -502,14 +503,14 @@ let ungen_upat lhs (sigma, uc, t) u =
let nb_cs_proj_args pc f u =
let na k =
List.length (snd (lookup_canonical_conversion (ConstRef pc, k))).o_TCOMPS in
- let nargs_of_proj t = match kind_of_term t with
+ let nargs_of_proj t = match kind t with
| App(_,args) -> Array.length args
| Proj _ -> 0 (* if splay_app calls expand_projection, this has to be
the number of arguments including the projected *)
| _ -> assert false in
- try match kind_of_term f with
+ try match kind f with
| Prod _ -> na Prod_cs
- | Sort s -> na (Sort_cs (family_of_sort s))
+ | Sort s -> na (Sort_cs (Sorts.family s))
| Const (c',_) when Constant.equal c' pc -> nargs_of_proj u.up_f
| Proj (c',_) when Constant.equal (Projection.constant c') pc -> nargs_of_proj u.up_f
| Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f))
@@ -517,22 +518,22 @@ let nb_cs_proj_args pc f u =
with Not_found -> -1
let isEvar_k k f =
- match kind_of_term f with Evar (k', _) -> k = k' | _ -> false
+ match kind f with Evar (k', _) -> k = k' | _ -> false
let nb_args c =
- match kind_of_term c with App (_, a) -> Array.length a | _ -> 0
+ match kind c with App (_, a) -> Array.length a | _ -> 0
let mkSubArg i a = if i = Array.length a then a else Array.sub a 0 i
let mkSubApp f i a = if i = 0 then f else mkApp (f, mkSubArg i a)
let splay_app ise =
- let rec loop c a = match kind_of_term c with
+ let rec loop c a = match kind c with
| App (f, a') -> loop f (Array.append a' a)
| Cast (c', _, _) -> loop c' a
| Evar ex ->
(try loop (existential_value ise ex) a with _ -> c, a)
| _ -> c, a in
- fun c -> match kind_of_term c with
+ fun c -> match kind c with
| App (f, a) -> loop f a
| Cast _ | Evar _ -> loop c [| |]
| _ -> c, [| |]
@@ -541,8 +542,8 @@ let filter_upat i0 f n u fpats =
let na = Array.length u.up_a in
if n < na then fpats else
let np = match u.up_k with
- | KpatConst when Term.eq_constr u.up_f f -> na
- | KpatFixed when Term.eq_constr u.up_f f -> na
+ | KpatConst when equal u.up_f f -> na
+ | KpatFixed when equal u.up_f f -> na
| KpatEvar k when isEvar_k k f -> na
| KpatLet when isLetIn f -> na
| KpatLam when isLambda f -> na
@@ -554,7 +555,7 @@ let filter_upat i0 f n u fpats =
if np < na then fpats else
let () = if !i0 < np then i0 := n in (u, np) :: fpats
-let eq_prim_proj c t = match kind_of_term t with
+let eq_prim_proj c t = match kind t with
| Proj(p,_) -> Constant.equal (Projection.constant p) c
| _ -> false
@@ -562,13 +563,13 @@ let filter_upat_FO i0 f n u fpats =
let np = nb_args u.up_FO in
if n < np then fpats else
let ok = match u.up_k with
- | KpatConst -> Term.eq_constr u.up_f f
- | KpatFixed -> Term.eq_constr u.up_f f
+ | KpatConst -> equal u.up_f f
+ | KpatFixed -> equal u.up_f f
| KpatEvar k -> isEvar_k k f
| KpatLet -> isLetIn f
| KpatLam -> isLambda f
| KpatRigid -> isRigid f
- | KpatProj pc -> Term.eq_constr f (mkConst pc) || eq_prim_proj pc f
+ | KpatProj pc -> equal f (mkConst pc) || eq_prim_proj pc f
| KpatFlex -> i0 := n; true in
if ok then begin if !i0 < np then i0 := np; (u, np) :: fpats end else fpats
@@ -741,13 +742,13 @@ let mk_tpattern_matcher ?(all_instances=false)
let x, pv, t, pb = destLetIn u.up_f in
let env' =
Environ.push_rel (Context.Rel.Declaration.LocalAssum(x, t)) env in
- let match_let f = match kind_of_term f with
+ let match_let f = match kind f with
| LetIn (_, v, _, b) -> unif_EQ env sigma pv v && unif_EQ env' sigma pb b
| _ -> false in match_let
- | KpatFixed -> Term.eq_constr u.up_f
- | KpatConst -> Term.eq_constr u.up_f
+ | KpatFixed -> equal u.up_f
+ | KpatConst -> equal u.up_f
| KpatLam -> fun c ->
- (match kind_of_term c with
+ (match kind c with
| Lambda _ -> unif_EQ env sigma u.up_f c
| _ -> false)
| _ -> unif_EQ env sigma u.up_f in
@@ -778,8 +779,8 @@ let rec uniquize = function
let t1 = nf_evar sigma1 t1 in
let f1 = nf_evar sigma1 f1 in
let a1 = Array.map (nf_evar sigma1) a1 in
- not (Term.eq_constr t t1 &&
- Term.eq_constr f f1 && CArray.for_all2 Term.eq_constr a a1) in
+ not (equal t t1 &&
+ equal f f1 && CArray.for_all2 equal a a1) in
x :: uniquize (List.filter neq xs) in
((fun env c h ~k ->
@@ -1018,7 +1019,7 @@ let input_ssrtermkind strm = match stream_nth 0 strm with
| Tok.KEYWORD "(" -> '('
| Tok.KEYWORD "@" -> '@'
| _ -> ' '
-let ssrtermkind = Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind
+let ssrtermkind = Pcoq.Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind
let interp_ssrterm _ gl t = Tacmach.project gl, t
@@ -1100,7 +1101,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
let decodeG t f g = decode ist (mkG t) f g in
let bad_enc id _ = CErrors.anomaly (str"bad encoding for pattern "++str id++str".") in
let cleanup_XinE h x rp sigma =
- let h_k = match kind_of_term h with Evar (k,_) -> k | _ -> assert false in
+ let h_k = match kind h with Evar (k,_) -> k | _ -> assert false in
let to_clean, update = (* handle rename if x is already used *)
let ctx = pf_hyps gl in
let len = Context.Named.length ctx in
@@ -1115,11 +1116,11 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
with Not_found -> ref (Some x), fun _ -> () in
let sigma0 = project gl in
let new_evars =
- let rec aux acc t = match kind_of_term t with
+ let rec aux acc t = match kind t with
| Evar (k,_) ->
if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else
(update k; k::acc)
- | _ -> fold_constr aux acc t in
+ | _ -> CoqConstr.fold aux acc t in
aux [] (nf_evar sigma rp) in
let sigma =
List.fold_left (fun sigma e ->
@@ -1202,7 +1203,7 @@ let interp_cpattern ist gl red redty = interp_pattern ist gl (T red) redty;;
let interp_rpattern ~wit_ssrpatternarg ist gl red = interp_pattern ~wit_ssrpatternarg ist gl red None;;
let id_of_pattern = function
- | _, T t -> (match kind_of_term t with Var id -> Some id | _ -> None)
+ | _, T t -> (match kind t with Var id -> Some id | _ -> None)
| _ -> None
(* The full occurrence set *)
@@ -1222,7 +1223,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
Evd.add (Evd.remove sigma e) e {e_def with Evd.evar_body = Evar_empty} in
sigma, e_body in
let ex_value hole =
- match kind_of_term hole with Evar (e,_) -> e | _ -> assert false in
+ match kind hole with Evar (e,_) -> e | _ -> assert false in
let mk_upat_for ?hack env sigma0 (sigma, t) ?(p=t) ok =
let sigma,pat= mk_tpattern ?hack env sigma0 (sigma,p) ok L2R (fs sigma t) in
sigma, [pat] in
@@ -1407,7 +1408,7 @@ let () =
let ssrinstancesof ist arg gl =
let ok rhs lhs ise = true in
-(* not (Term.eq_constr lhs (Evarutil.nf_evar ise rhs)) in *)
+(* not (equal lhs (Evarutil.nf_evar ise rhs)) in *)
let env, sigma, concl = pf_env gl, project gl, pf_concl gl in
let concl = EConstr.Unsafe.to_constr concl in
let sigma0, cpat = interp_cpattern ist gl arg None in
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
index 8e2a1a717..8ab666f7e 100644
--- a/plugins/ssrmatching/ssrmatching.mli
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -6,7 +6,7 @@ open Genarg
open Tacexpr
open Environ
open Evd
-open Term
+open Constr
(** ******** Small Scale Reflection pattern matching facilities ************* *)
diff --git a/pretyping/arguments_renaming.ml b/pretyping/arguments_renaming.ml
index ea33f1c0d..d59102b6c 100644
--- a/pretyping/arguments_renaming.ml
+++ b/pretyping/arguments_renaming.ml
@@ -10,6 +10,7 @@
open Names
open Globnames
open Term
+open Constr
open Environ
open Util
open Libobject
@@ -103,7 +104,7 @@ let rename_type_of_constructor env cstruct =
let rename_typing env c =
let j = Typeops.infer env c in
let j' =
- match kind_of_term c with
+ match kind c with
| Const (c,u) -> { j with uj_type = rename_type j.uj_type (ConstRef c) }
| Ind (i,u) -> { j with uj_type = rename_type j.uj_type (IndRef i) }
| Construct (k,u) -> { j with uj_type = rename_type j.uj_type (ConstructRef k) }
diff --git a/pretyping/arguments_renaming.mli b/pretyping/arguments_renaming.mli
index 804e38216..b499da3ab 100644
--- a/pretyping/arguments_renaming.mli
+++ b/pretyping/arguments_renaming.mli
@@ -9,7 +9,7 @@
open Names
open Globnames
open Environ
-open Term
+open Constr
val rename_arguments : bool -> global_reference -> Name.t list -> unit
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index aefa09dbe..4f3669a2b 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -13,7 +13,7 @@ open CErrors
open Util
open Names
open Nameops
-open Term
+open Constr
open Termops
open Environ
open EConstr
@@ -1014,7 +1014,7 @@ let adjust_impossible_cases pb pred tomatch submat =
this means that the Evd.define below may redefine an already defined
evar. See e.g. first definition of test for bug #3388. *)
let pred = EConstr.Unsafe.to_constr pred in
- begin match kind_of_term pred with
+ begin match Constr.kind pred with
| Evar (evk,_) when snd (evar_source evk !(pb.evdref)) == Evar_kinds.ImpossibleCase ->
if not (Evd.is_defined !(pb.evdref) evk) then begin
let evd, default = use_unit_judge !(pb.evdref) in
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 327ee0015..3a139b7b0 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Evd
open Environ
open EConstr
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 19d61a64d..3a2eac7e7 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -8,7 +8,7 @@
open Util
open Names
-open Term
+open Constr
open Vars
open CClosure
open Esubst
@@ -211,7 +211,7 @@ and reify_value = function (* reduction under binders *)
t
(* map_constr_with_binders subs_lift (cbv_norm_term) env t *)
| LAM (n,ctxt,b,env) ->
- List.fold_left (fun c (n,t) -> Term.mkLambda (n, t, c)) b ctxt
+ List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) b ctxt
| FIXP ((lij,(names,lty,bds)),env,args) ->
mkApp
(mkFix (lij,
@@ -240,7 +240,7 @@ and reify_value = function (* reduction under binders *)
let rec norm_head info env t stack =
(* no reduction under binders *)
- match kind_of_term t with
+ match kind t with
(* stack grows (remove casts) *)
| App (head,args) -> (* Applied terms are normalized immediately;
they could be computed when getting out of the stack *)
@@ -294,7 +294,7 @@ let rec norm_head info env t stack =
(* non-neutral cases *)
| Lambda _ ->
- let ctxt,b = decompose_lam t in
+ let ctxt,b = Term.decompose_lam t in
(LAM(List.length ctxt, List.rev ctxt,b,env), stack)
| Fix fix -> (FIXP(fix,env,[||]), stack)
| CoFix cofix -> (COFIXP(cofix,env,[||]), stack)
@@ -411,12 +411,12 @@ and cbv_norm_value info = function (* reduction under binders *)
| STACK (n,v,stk) ->
lift n (apply_stack info (cbv_norm_value info v) stk)
| CBN(t,env) ->
- map_constr_with_binders subs_lift (cbv_norm_term info) env t
+ Constr.map_with_binders subs_lift (cbv_norm_term info) env t
| LAM (n,ctxt,b,env) ->
let nctxt =
List.map_i (fun i (x,ty) ->
(x,cbv_norm_term info (subs_liftn i env) ty)) 0 ctxt in
- compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b)
+ Term.compose_lam (List.rev nctxt) (cbv_norm_term info (subs_liftn n env) b)
| FIXP ((lij,(names,lty,bds)),env,args) ->
mkApp
(mkFix (lij,
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index 3ee7bebf0..5f9609a5c 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -24,7 +24,7 @@ val cbv_norm : cbv_infos -> constr -> constr
(***********************************************************************
i This is for cbv debug *)
-open Term
+open Constr
type cbv_value =
| VAL of int * constr
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index fe969f9e5..c36630f5d 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -10,12 +10,12 @@ open CErrors
open Util
open Pp
open Names
+open Constr
open Libnames
open Globnames
open Nametab
open Environ
open Libobject
-open Term
open Mod_subst
(* usage qque peu general: utilise aussi dans record *)
@@ -43,7 +43,7 @@ type coe_info_typ = {
coe_value : constr;
coe_type : types;
coe_local : bool;
- coe_context : Univ.universe_context_set;
+ coe_context : Univ.ContextSet.t;
coe_is_identity : bool;
coe_is_projection : bool;
coe_param : int }
diff --git a/pretyping/constr_matching.mli b/pretyping/constr_matching.mli
index 34c62043e..780ccc23d 100644
--- a/pretyping/constr_matching.mli
+++ b/pretyping/constr_matching.mli
@@ -9,7 +9,7 @@
(** This module implements pattern-matching on terms *)
open Names
-open Term
+open Constr
open EConstr
open Environ
open Pattern
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index c02fc5aaf..0d1e401d9 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -554,7 +554,7 @@ and detype_r d flags avoid env sigma t =
let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel sigma c && Int.Set.mem (destRel sigma c) rels || isVar sigma c && (Id.Set.mem (destVar sigma c) fvs)))) (Evd.find sigma evk) cl in
id,l
with Not_found ->
- Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
+ Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
(Array.map_to_list (fun c -> (Id.of_string "__",c)) cl)
in
GEvar (id,
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index f03bde68e..cb1c0d8d4 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Names
-open Term
open Environ
open EConstr
open Glob_term
@@ -40,7 +39,7 @@ val detype_names : bool -> Id.Set.t -> names_context -> env -> evar_map -> const
val detype : 'a delay -> ?lax:bool -> bool -> Id.Set.t -> env -> evar_map -> constr -> 'a glob_constr_g
-val detype_sort : evar_map -> sorts -> glob_sort
+val detype_sort : evar_map -> Sorts.t -> glob_sort
val detype_rel_context : 'a delay -> ?lax:bool -> constr option -> Id.Set.t -> (names_context * env) ->
evar_map -> rel_context -> 'a glob_decl_g list
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 45fe2b2d8..681eb17d3 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -10,6 +10,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Termops
open Environ
open EConstr
@@ -49,7 +50,7 @@ let _ = Goptions.declare_bool_option {
let impossible_default_case () =
let c, ctx = Universes.fresh_global_instance (Global.env()) (Globnames.ConstRef Coqlib.id) in
let (_, u) = Term.destConst c in
- Some (c, Term.mkConstU (Coqlib.type_of_id, u), ctx)
+ Some (c, Constr.mkConstU (Coqlib.type_of_id, u), ctx)
let coq_unit_judge =
let open Environ in
@@ -175,7 +176,7 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
| Sort s ->
let s = ESorts.kind sigma s in
lookup_canonical_conversion
- (proji, Sort_cs (family_of_sort s)),[]
+ (proji, Sort_cs (Sorts.family s)),[]
| Proj (p, c) ->
let c2 = Globnames.ConstRef (Projection.constant p) in
let c = Retyping.expand_projection env sigma p c [] in
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index c30d1d26b..d793b06d3 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -47,7 +47,7 @@ val check_problems_are_solved : env -> evar_map -> unit
val check_conv_record : env -> evar_map ->
state -> state ->
- Univ.universe_context_set * (constr * constr)
+ Univ.ContextSet.t * (constr * constr)
* constr * constr list * (constr Stack.t * constr Stack.t) *
(constr Stack.t * constr Stack.t) *
(constr Stack.t * constr Stack.t) * constr *
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index 5f12f360b..18dbbea1b 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -10,6 +10,7 @@ open Util
open Pp
open Names
open Term
+open Constr
open Termops
open EConstr
open Vars
@@ -82,7 +83,7 @@ let define_pure_evar_as_product evd evk =
let newenv = push_named (LocalAssum (id, dom)) evenv in
let src = evar_source evk evd1 in
let filter = Filter.extend 1 (evar_filter evi) in
- if is_prop_sort (ESorts.kind evd1 s) then
+ if Sorts.is_prop (ESorts.kind evd1 s) then
(* Impredicative product, conclusion must fall in [Prop]. *)
new_evar newenv evd1 concl ~src ~filter
else
diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli
index 5477c5c99..869e3adbf 100644
--- a/pretyping/evardefine.mli
+++ b/pretyping/evardefine.mli
@@ -7,7 +7,6 @@
(************************************************************************)
open Names
-open Term
open EConstr
open Evd
open Environ
@@ -39,7 +38,7 @@ val lift_tycon : int -> type_constraint -> type_constraint
val define_evar_as_product : evar_map -> existential -> evar_map * types
val define_evar_as_lambda : env -> evar_map -> existential -> evar_map * types
-val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
+val define_evar_as_sort : env -> evar_map -> existential -> evar_map * Sorts.t
(** {6 debug pretty-printer:} *)
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index b906c3b59..fba154291 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -10,6 +10,7 @@ open Util
open CErrors
open Names
open Term
+open Constr
open Environ
open Termops
open Evd
@@ -1391,7 +1392,7 @@ let occur_evar_upto_types sigma n c =
let c = EConstr.Unsafe.to_constr c in
let seen = ref Evar.Set.empty in
(** FIXME: Is that supposed to be evar-insensitive? *)
- let rec occur_rec c = match kind_of_term c with
+ let rec occur_rec c = match Constr.kind c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar (sp,args as e) ->
if Evar.Set.mem sp !seen then
diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli
index 811b4dc18..e5d288b5c 100644
--- a/pretyping/evarsolve.mli
+++ b/pretyping/evarsolve.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open EConstr
open Evd
open Environ
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 0cb5974e8..48b33e708 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -18,6 +18,7 @@ open Libnames
open Globnames
open Nameops
open Term
+open Constr
open Vars
open Namegen
open Declarations
@@ -33,7 +34,7 @@ type dep_flag = bool
(* Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * Sorts.t * pinductive
| NotMutualInScheme of inductive * inductive
| NotAllowedDependentAnalysis of (*isrec:*) bool * inductive
@@ -168,7 +169,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in
let p' = EConstr.Unsafe.to_constr p' in
let largs = List.map EConstr.Unsafe.to_constr largs in
- match kind_of_term p' with
+ match kind p' with
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
make_prod env (n,t,prec (push_rel d env) (i+1) (d::sign) c)
@@ -186,13 +187,13 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
| _ ->
let t' = whd_all env sigma (EConstr.of_constr p) in
let t' = EConstr.Unsafe.to_constr t' in
- if Term.eq_constr p' t' then assert false
+ if Constr.equal p' t' then assert false
else prec env i sign t'
in
prec env 0 []
in
let rec process_constr env i c recargs nhyps li =
- if nhyps > 0 then match kind_of_term c with
+ if nhyps > 0 then match kind c with
| Prod (n,t,c_0) ->
let (optionpos,rest) =
match recargs with
@@ -247,7 +248,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let p',largs = whd_allnolet_stack env sigma (EConstr.of_constr p) in
let p' = EConstr.Unsafe.to_constr p' in
let largs = List.map EConstr.Unsafe.to_constr largs in
- match kind_of_term p' with
+ match kind p' with
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
mkLambda_name env (n,t,prec (push_rel d env) (i+1) (d::hyps) c)
@@ -261,7 +262,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
| _ ->
let t' = whd_all env sigma (EConstr.of_constr p) in
let t' = EConstr.Unsafe.to_constr t' in
- if Term.eq_constr t' p' then assert false
+ if Constr.equal t' p' then assert false
else prec env i hyps t'
in
prec env 0 []
@@ -505,7 +506,7 @@ let build_case_analysis_scheme_default env sigma pity kind =
[rec] by [s] *)
let change_sort_arity sort =
- let rec drec a = match kind_of_term a with
+ let rec drec a = match kind a with
| Cast (c,_,_) -> drec c
| Prod (n,t,c) -> let s, c' = drec c in s, mkProd (n, t, c')
| LetIn (n,b,t,c) -> let s, c' = drec c in s, mkLetIn (n,b,t,c')
@@ -519,7 +520,7 @@ let change_sort_arity sort =
let weaken_sort_scheme env evd set sort npars term ty =
let evdref = ref evd in
let rec drec np elim =
- match kind_of_term elim with
+ match kind elim with
| Prod (n,t,c) ->
if Int.equal np 0 then
let osort, t' = change_sort_arity sort t in
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index 2825c4d83..a9838cffe 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -7,14 +7,14 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Evd
(** Errors related to recursors building *)
type recursion_scheme_error =
- | NotAllowedCaseAnalysis of (*isrec:*) bool * sorts * pinductive
+ | NotAllowedCaseAnalysis of (*isrec:*) bool * Sorts.t * pinductive
| NotMutualInScheme of inductive * inductive
| NotAllowedDependentAnalysis of (*isrec:*) bool * inductive
@@ -27,25 +27,25 @@ type dep_flag = bool
(** Build a case analysis elimination scheme in some sort family *)
val build_case_analysis_scheme : env -> Evd.evar_map -> pinductive ->
- dep_flag -> sorts_family -> evar_map * Constr.t
+ dep_flag -> Sorts.family -> evar_map * Constr.t
(** Build a dependent case elimination predicate unless type is in Prop
or is a recursive record with primitive projections. *)
val build_case_analysis_scheme_default : env -> evar_map -> pinductive ->
- sorts_family -> evar_map * Constr.t
+ Sorts.family -> evar_map * Constr.t
(** Builds a recursive induction scheme (Peano-induction style) in the same
sort family as the inductive family; it is dependent if not in Prop
or a recursive record with primitive projections. *)
val build_induction_scheme : env -> evar_map -> pinductive ->
- dep_flag -> sorts_family -> evar_map * constr
+ dep_flag -> Sorts.family -> evar_map * constr
(** Builds mutual (recursive) induction schemes *)
val build_mutual_induction_scheme :
- env -> evar_map -> (pinductive * dep_flag * sorts_family) list -> evar_map * constr list
+ env -> evar_map -> (pinductive * dep_flag * Sorts.family) list -> evar_map * constr list
(** Scheme combinators *)
@@ -54,13 +54,13 @@ val build_mutual_induction_scheme :
scheme quantified on sort [s]. [set] asks for [s] be declared equal to [i],
otherwise just less or equal to [i]. *)
-val weaken_sort_scheme : env -> evar_map -> bool -> sorts -> int -> constr -> types ->
+val weaken_sort_scheme : env -> evar_map -> bool -> Sorts.t -> int -> constr -> types ->
evar_map * types * constr
(** Recursor names utilities *)
-val lookup_eliminator : inductive -> sorts_family -> Globnames.global_reference
-val elimination_suffix : sorts_family -> string
-val make_elimination_ident : Id.t -> sorts_family -> Id.t
+val lookup_eliminator : inductive -> Sorts.family -> Globnames.global_reference
+val elimination_suffix : Sorts.family -> string
+val make_elimination_ident : Id.t -> Sorts.family -> Id.t
val case_suffix : string
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index b31ee03d8..34df7d3d7 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -11,6 +11,7 @@ open Util
open Names
open Univ
open Term
+open Constr
open Vars
open Termops
open Declarations
@@ -643,7 +644,7 @@ let type_of_projection_knowing_arg env sigma p c ty =
syntactic conditions *)
let control_only_guard env c =
- let check_fix_cofix e c = match kind_of_term c with
+ let check_fix_cofix e c = match kind c with
| CoFix (_,(_,_,_) as cofix) ->
Inductive.check_cofix e cofix
| Fix (_,(_,_,_) as fix) ->
@@ -659,7 +660,7 @@ let control_only_guard env c =
(* inference of subtyping condition for inductive types *)
let infer_inductive_subtyping_arity_constructor
- (env, evd, csts) (subst : constr -> constr) (arcn : Term.types) is_arity (params : Context.Rel.t) =
+ (env, evd, csts) (subst : constr -> constr) (arcn : types) is_arity (params : Context.Rel.t) =
let numchecked = ref 0 in
let numparams = Context.Rel.nhyps params in
let update_contexts (env, evd, csts) csts' =
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 65d5161df..febe99b0b 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Declarations
open Environ
open Evd
@@ -120,7 +120,7 @@ val constructor_nrealdecls_env : env -> constructor -> int
val constructor_has_local_defs : constructor -> bool
val inductive_has_local_defs : inductive -> bool
-val allowed_sorts : env -> inductive -> sorts_family list
+val allowed_sorts : env -> inductive -> Sorts.family list
(** (Co)Inductive records with primitive projections do not have eta-conversion,
hence no dependent elimination. *)
@@ -152,12 +152,12 @@ val get_projections : env -> inductive_family -> Constant.t array option
(** [get_arity] returns the arity of the inductive family instantiated
with the parameters; if recursively non-uniform parameters are not
part of the inductive family, they appears in the arity *)
-val get_arity : env -> inductive_family -> Context.Rel.t * sorts_family
+val get_arity : env -> inductive_family -> Context.Rel.t * Sorts.family
val build_dependent_constructor : constructor_summary -> constr
val build_dependent_inductive : env -> inductive_family -> constr
val make_arity_signature : env -> evar_map -> bool -> inductive_family -> EConstr.rel_context
-val make_arity : env -> evar_map -> bool -> inductive_family -> sorts -> EConstr.types
+val make_arity : env -> evar_map -> bool -> inductive_family -> Sorts.t -> EConstr.types
val build_branch_type : env -> evar_map -> bool -> constr -> constructor_summary -> types
(** Raise [Not_found] if not given a valid inductive type *)
@@ -172,7 +172,7 @@ val find_coinductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.
(** Builds the case predicate arity (dependent or not) *)
val arity_of_case_predicate :
- env -> inductive_family -> bool -> sorts -> types
+ env -> inductive_family -> bool -> Sorts.t -> types
val type_case_branches_with_names :
env -> evar_map -> pinductive * EConstr.constr list -> constr -> constr -> types array * types
@@ -203,8 +203,8 @@ val control_only_guard : env -> types -> unit
(* inference of subtyping condition for inductive types *)
(* for debugging purposes only to be removed *)
val infer_inductive_subtyping_arity_constructor : Environ.env * Evd.evar_map * Univ.Constraint.t ->
-(Term.constr -> Term.constr) ->
-Term.types -> bool -> Context.Rel.t -> Environ.env * Evd.evar_map * Univ.Constraint.t
+(constr -> constr) ->
+types -> bool -> Context.Rel.t -> Environ.env * Evd.evar_map * Univ.Constraint.t
val infer_inductive_subtyping : Environ.env -> Evd.evar_map -> Entries.mutual_inductive_entry ->
Entries.mutual_inductive_entry
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index fe134f512..dafe8cb26 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -7,6 +7,7 @@
(************************************************************************)
open CErrors
open Term
+open Constr
open Vars
open Environ
open Reduction
@@ -98,7 +99,7 @@ let app_type env c =
let find_rectype_a env c =
let (t, l) = app_type env c in
- match kind_of_term t with
+ match kind t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
@@ -131,7 +132,7 @@ let construct_of_constr_notnative const env tag (mind, _ as ind) u allargs =
let construct_of_constr const env tag typ =
let t, l = app_type env typ in
- match kind_of_term t with
+ match kind t with
| Ind (ind,u) ->
construct_of_constr_notnative const env tag ind u l
| _ -> assert false
@@ -360,7 +361,7 @@ and nf_atom_type env sigma atom =
and nf_predicate env sigma ind mip params v pT =
- match kind_of_value v, kind_of_term pT with
+ match kind_of_value v, kind pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
let vb = f (mk_rel_accu k) in
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index fd76a9473..ee79b5474 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -12,6 +12,7 @@ open Names
open Globnames
open Nameops
open Term
+open Constr
open Vars
open Glob_term
open Pp
@@ -75,8 +76,8 @@ and cofixpoint_eq (i1, r1) (i2, r2) =
and rec_declaration_eq (n1, c1, r1) (n2, c2, r2) =
Array.equal Name.equal n1 n2 &&
- Array.equal Term.eq_constr c1 c2 &&
- Array.equal Term.eq_constr r1 r2
+ Array.equal Constr.equal c1 c2 &&
+ Array.equal Constr.equal r1 r2
let rec occur_meta_pattern = function
| PApp (f,args) ->
@@ -149,7 +150,7 @@ let head_of_constr_reference sigma c = match EConstr.kind sigma c with
let pattern_of_constr env sigma t =
let rec pattern_of_constr env t =
let open Context.Rel.Declaration in
- match kind_of_term t with
+ match kind t with
| Rel n -> PRel n
| Meta n -> PMeta (Some (Id.of_string ("META" ^ string_of_int n)))
| Var id -> PVar id
@@ -165,7 +166,7 @@ let pattern_of_constr env sigma t =
pattern_of_constr (push_rel (LocalAssum (na, c)) env) b)
| App (f,a) ->
(match
- match kind_of_term f with
+ match kind f with
| Evar (evk,args) ->
(match snd (Evd.evar_source evk sigma) with
Evar_kinds.MatchingVar (Evar_kinds.SecondOrderPatVar id) -> Some id
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 54b477bed..ce478ac20 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open EConstr
open Type_errors
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 124fa6e06..dab376ef0 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open EConstr
open Type_errors
@@ -99,8 +99,8 @@ val error_ill_typed_rec_body :
val error_elim_arity :
?loc:Loc.t -> env -> Evd.evar_map ->
- pinductive -> sorts_family list -> constr ->
- unsafe_judgment -> (sorts_family * sorts_family * arity_error) option -> 'b
+ pinductive -> Sorts.family list -> constr ->
+ unsafe_judgment -> (Sorts.family * Sorts.family * arity_error) option -> 'b
val error_not_a_type :
?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 6537d1ecf..eb2b435bf 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -12,7 +12,7 @@
into elementary ones, insertion of coercions and resolution of
implicit arguments. *)
-open Term
+open Constr
open Environ
open Evd
open EConstr
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 4f4e49d82..cb24ca804 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -19,7 +19,7 @@ open Pp
open Names
open Globnames
open Nametab
-open Term
+open Constr
open Libobject
open Mod_subst
open Reductionops
@@ -144,7 +144,7 @@ type obj_typ = {
type cs_pattern =
Const_cs of global_reference
| Prod_cs
- | Sort_cs of sorts_family
+ | Sort_cs of Sorts.family
| Default_cs
let eq_cs_pattern p1 p2 = match p1, p2 with
@@ -172,7 +172,7 @@ let keep_true_projections projs kinds =
List.map_filter filter (List.combine projs kinds)
let cs_pattern_of_constr env t =
- match kind_of_term t with
+ match kind t with
App (f,vargs) ->
begin
try Const_cs (global_of_constr f) , None, Array.to_list vargs
@@ -184,7 +184,7 @@ let cs_pattern_of_constr env t =
let { Environ.uj_type = ty } = Typeops.infer env c in
let _, params = Inductive.find_rectype env ty in
Const_cs (ConstRef (Projection.constant p)), None, params @ [c]
- | Sort s -> Sort_cs (family_of_sort s), None, []
+ | Sort s -> Sort_cs (Sorts.family s), None, []
| _ ->
begin
try Const_cs (global_of_constr t) , None, []
@@ -213,7 +213,7 @@ let compute_canonical_projections warn (con,ind) =
let sign = List.map (on_snd EConstr.Unsafe.to_constr) sign in
let t = EConstr.Unsafe.to_constr t in
let lt = List.rev_map snd sign in
- let args = snd (decompose_app t) in
+ let args = snd (Term.decompose_app t) in
let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
let params, projs = List.chop p args in
@@ -311,10 +311,10 @@ let check_and_decompose_canonical_structure ref =
| None -> error_not_structure ref in
let body = snd (splay_lam (Global.env()) Evd.empty (EConstr.of_constr vc)) (** FIXME *) in
let body = EConstr.Unsafe.to_constr body in
- let f,args = match kind_of_term body with
+ let f,args = match kind body with
| App (f,args) -> f,args
| _ -> error_not_structure ref in
- let indsp = match kind_of_term f with
+ let indsp = match kind f with
| Construct ((indsp,1),u) -> indsp
| _ -> error_not_structure ref in
let s = try lookup_structure indsp with Not_found -> error_not_structure ref in
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index 066623164..f15418577 100644
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Globnames
(** Operations concerning records and canonical structures *)
@@ -52,7 +52,7 @@ val find_projection : global_reference -> struc_typ
type cs_pattern =
Const_cs of global_reference
| Prod_cs
- | Sort_cs of sorts_family
+ | Sort_cs of Sorts.family
| Default_cs
type obj_typ = {
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 308896f3a..a87debeff 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -9,7 +9,7 @@
open CErrors
open Util
open Names
-open Term
+open Constr
open Termops
open Univ
open Evd
@@ -339,7 +339,7 @@ struct
type 'a member =
| App of 'a app_node
- | Case of Term.case_info * 'a * 'a array * Cst_stack.t
+ | Case of case_info * 'a * 'a array * Cst_stack.t
| Proj of int * int * projection * Cst_stack.t
| Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t
| Cst of cst_member * int * int list * 'a t * Cst_stack.t
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 950fcdee4..db0c29aef 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open EConstr
open Univ
open Evd
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 56f8b33e0..5dd6879d3 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -10,6 +10,7 @@ open Pp
open CErrors
open Util
open Term
+open Constr
open Inductive
open Inductiveops
open Names
@@ -146,7 +147,7 @@ let retype ?(polyprop=true) sigma =
| Cast (c,_, s) when isSort sigma s -> destSort sigma s
| Sort s ->
begin match ESorts.kind sigma s with
- | Prop _ -> type1_sort
+ | Prop _ -> Sorts.type1
| Type u -> Type (Univ.super u)
end
| Prod (name,t,c2) ->
@@ -167,7 +168,7 @@ let retype ?(polyprop=true) sigma =
and sort_family_of env t =
match EConstr.kind sigma t with
- | Cast (c,_, s) when isSort sigma s -> family_of_sort (destSort sigma s)
+ | Cast (c,_, s) when isSort sigma s -> Sorts.family (destSort sigma s)
| Sort _ -> InType
| Prod (name,t,c2) ->
let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in
@@ -175,12 +176,12 @@ let retype ?(polyprop=true) sigma =
s2 == InSet && sort_family_of env t == InType then InType else s2
| App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
- family_of_sort (sort_of_atomic_type env sigma t args)
+ Sorts.family (sort_of_atomic_type env sigma t args)
| App(f,args) ->
- family_of_sort (sort_of_atomic_type env sigma (type_of env f) args)
+ Sorts.family (sort_of_atomic_type env sigma (type_of env f) args)
| Lambda _ | Fix _ | Construct _ -> retype_error NotAType
| _ ->
- family_of_sort (decomp_sort env sigma (type_of env t))
+ Sorts.family (decomp_sort env sigma (type_of env t))
and type_of_global_reference_knowing_parameters env c args =
let argtyps =
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index ed3a9d0f9..af86df499 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -6,7 +6,6 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
open Evd
open Environ
open EConstr
@@ -30,10 +29,10 @@ val get_type_of :
?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> types
val get_sort_of :
- ?polyprop:bool -> env -> evar_map -> types -> sorts
+ ?polyprop:bool -> env -> evar_map -> types -> Sorts.t
val get_sort_family_of :
- ?polyprop:bool -> env -> evar_map -> types -> sorts_family
+ ?polyprop:bool -> env -> evar_map -> types -> Sorts.family
(** Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
@@ -44,7 +43,7 @@ val type_of_global_reference_knowing_parameters : env -> evar_map -> constr ->
val type_of_global_reference_knowing_conclusion :
env -> evar_map -> constr -> types -> evar_map * types
-val sorts_of_context : env -> evar_map -> rel_context -> sorts list
+val sorts_of_context : env -> evar_map -> rel_context -> Sorts.t list
val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 1e06e9f8c..d55b286fb 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -11,6 +11,7 @@ open Names
open Globnames
open Decl_kinds
open Term
+open Constr
open Vars
open Evd
open Util
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 20fbbb5a0..062d5cf35 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -8,7 +8,7 @@
open Names
open Globnames
-open Term
+open Constr
open Evd
open Environ
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index 1f35fa19a..43066c809 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -160,7 +160,7 @@ let check_type_fixpoint ?loc env evdref lna lar vdefj =
(* FIXME: might depend on the level of actual parameters!*)
let check_allowed_sort env sigma ind c p =
let pj = Retyping.get_judgment_of env sigma p in
- let ksort = family_of_sort (ESorts.kind sigma (sort_of_arity env sigma pj.uj_type)) in
+ let ksort = Sorts.family (ESorts.kind sigma (sort_of_arity env sigma pj.uj_type)) in
let specif = Global.lookup_inductive (fst ind) in
let sorts = elim_sorts specif in
if not (List.exists ((==) ksort) sorts) then
@@ -195,11 +195,11 @@ let check_cofix env sigma pcofix =
let judge_of_prop =
{ uj_val = EConstr.mkProp;
- uj_type = EConstr.mkSort type1_sort }
+ uj_type = EConstr.mkSort Sorts.type1 }
let judge_of_set =
{ uj_val = EConstr.mkSet;
- uj_type = EConstr.mkSort type1_sort }
+ uj_type = EConstr.mkSort Sorts.type1 }
let judge_of_prop_contents = function
| Null -> judge_of_prop
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 1e2078826..9f084ae8d 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open EConstr
open Evd
@@ -26,7 +26,7 @@ val type_of : ?refresh:bool -> env -> evar_map -> constr -> evar_map * types
val e_type_of : ?refresh:bool -> env -> evar_map ref -> constr -> types
(** Typecheck a type and return its sort *)
-val e_sort_of : env -> evar_map ref -> types -> sorts
+val e_sort_of : env -> evar_map ref -> types -> Sorts.t
(** Typecheck a term has a given type (assuming the type is OK) *)
val e_check : env -> evar_map ref -> constr -> types -> unit
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 01c28b5ed..be9943f33 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -12,7 +12,7 @@ open CErrors
open Pp
open Util
open Names
-open Term
+open Constr
open Termops
open Environ
open EConstr
@@ -68,10 +68,10 @@ let _ = Goptions.declare_bool_option {
let unsafe_occur_meta_or_existential c =
let c = EConstr.Unsafe.to_constr c in
- let rec occrec c = match kind_of_term c with
+ let rec occrec c = match Constr.kind c with
| Evar _ -> raise Occur
| Meta _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> Constr.iter occrec c
in try occrec c; false with Occur -> true
@@ -79,7 +79,7 @@ let occur_meta_or_undefined_evar evd c =
(** This is performance-critical. Using the evar-insensitive API changes the
resulting heuristic. *)
let c = EConstr.Unsafe.to_constr c in
- let rec occrec c = match kind_of_term c with
+ let rec occrec c = match Constr.kind c with
| Meta _ -> raise Occur
| Evar (ev,args) ->
(match evar_body (Evd.find evd ev) with
@@ -609,7 +609,7 @@ let subst_defined_metas_evars sigma (bl,el) c =
(** This seems to be performance-critical, and using the evar-insensitive
primitives blow up the time passed in this function. *)
let c = EConstr.Unsafe.to_constr c in
- let rec substrec c = match kind_of_term c with
+ let rec substrec c = match Constr.kind c with
| Meta i ->
let select (j,_,_) = Int.equal i j in
substrec (EConstr.Unsafe.to_constr (pi2 (List.find select bl)))
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index fce17d564..085e8c5b8 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open EConstr
open Environ
open Evd
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index 66cc42cb6..b5b8987e3 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -10,6 +10,7 @@ open Util
open Names
open Declarations
open Term
+open Constr
open Vars
open Environ
open Inductive
@@ -51,7 +52,7 @@ let invert_tag cst tag reloc_tbl =
let find_rectype_a env c =
let (t, l) = decompose_appvect (whd_all env c) in
- match kind_of_term t with
+ match kind t with
| Ind ind -> (ind, l)
| _ -> assert false
@@ -262,7 +263,7 @@ and nf_stk ?from:(from=0) env sigma c t stk =
nf_stk env sigma (mkProj(p',c)) ty stk
and nf_predicate env sigma ind mip params v pT =
- match whd_val v, kind_of_term pT with
+ match whd_val v, kind pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
let vb = body_of_vfun k f in
diff --git a/printing/printer.ml b/printing/printer.ml
index e4bb89304..751e91cf0 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -11,6 +11,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Environ
open Globnames
open Nametab
diff --git a/printing/printer.mli b/printing/printer.mli
index e74f47efe..cd6e6d63a 100644
--- a/printing/printer.mli
+++ b/printing/printer.mli
@@ -8,7 +8,7 @@
open Names
open Globnames
-open Term
+open Constr
open Environ
open Pattern
open Evd
@@ -95,15 +95,15 @@ val pr_constr_pattern : constr_pattern -> Pp.t
val pr_cases_pattern : cases_pattern -> Pp.t
-val pr_sort : evar_map -> sorts -> Pp.t
+val pr_sort : evar_map -> Sorts.t -> Pp.t
(** Universe constraints *)
val pr_polymorphic : bool -> Pp.t
val pr_cumulative : bool -> bool -> Pp.t
-val pr_universe_instance : evar_map -> Univ.universe_context -> Pp.t
-val pr_universe_ctx : evar_map -> Univ.universe_context -> Pp.t
-val pr_cumulativity_info : evar_map -> Univ.cumulativity_info -> Pp.t
+val pr_universe_instance : evar_map -> Univ.UContext.t -> Pp.t
+val pr_universe_ctx : evar_map -> Univ.UContext.t -> Pp.t
+val pr_cumulativity_info : evar_map -> Univ.CumulativityInfo.t -> Pp.t
(** Printing global references using names as short as possible *)
@@ -198,7 +198,7 @@ module ContextObjectMap : CMap.ExtS
with type key = context_object and module Set := ContextObjectSet
val pr_assumptionset :
- env -> Term.types ContextObjectMap.t -> Pp.t
+ env -> types ContextObjectMap.t -> Pp.t
val pr_goal_by_id : Id.t -> Pp.t
diff --git a/printing/printmod.ml b/printing/printmod.ml
index d4b756346..6b3b177aa 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -7,7 +7,7 @@
(************************************************************************)
open Util
-open Term
+open Constr
open Pp
open Names
open Environ
@@ -149,7 +149,7 @@ let print_mutual_inductive env mind mib =
let get_fields =
let rec prodec_rec l subst c =
- match kind_of_term c with
+ match kind c with
| Prod (na,t,c) ->
let id = match na with Name id -> id | Anonymous -> Id.of_string "_" in
prodec_rec ((id,true,Vars.substl subst t)::l) (mkVar id::subst) c
diff --git a/proofs/clenv.mli b/proofs/clenv.mli
index 9c69995f4..9a2026dd3 100644
--- a/proofs/clenv.mli
+++ b/proofs/clenv.mli
@@ -11,7 +11,7 @@
evar-based clauses. *)
open Names
-open Term
+open Constr
open Environ
open Evd
open EConstr
diff --git a/proofs/goal.ml b/proofs/goal.ml
index 61f3e4a02..19f816a01 100644
--- a/proofs/goal.ml
+++ b/proofs/goal.ml
@@ -99,7 +99,7 @@ module V82 = struct
let same_goal evars1 gl1 evars2 gl2 =
let evi1 = Evd.find evars1 gl1 in
let evi2 = Evd.find evars2 gl2 in
- Term.eq_constr evi1.Evd.evar_concl evi2.Evd.evar_concl &&
+ Constr.equal evi1.Evd.evar_concl evi2.Evd.evar_concl &&
Environ.eq_named_context_val evi1.Evd.evar_hyps evi2.Evd.evar_hyps
let weak_progress glss gls =
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 20d075ae1..a633238f4 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -12,6 +12,7 @@ open Util
open Names
open Nameops
open Term
+open Constr
open Vars
open Termops
open Environ
@@ -284,12 +285,12 @@ let error_unsupported_deep_meta c =
strbrk "supported; try \"refine\" instead.")
let collect_meta_variables c =
- let rec collrec deep acc c = match kind_of_term c with
+ let rec collrec deep acc c = match kind c with
| Meta mv -> if deep then error_unsupported_deep_meta () else mv::acc
| Cast(c,_,_) -> collrec deep acc c
- | (App _| Case _) -> Term.fold_constr (collrec deep) acc c
+ | (App _| Case _) -> Constr.fold (collrec deep) acc c
| Proj (_, c) -> collrec deep acc c
- | _ -> Term.fold_constr (collrec true) acc c
+ | _ -> Constr.fold (collrec true) acc c
in
List.rev (collrec false [] c)
@@ -332,7 +333,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
let sigma = check_conv_leq_goal env sigma trm t'ty conclty in
(goalacc,t'ty,sigma,trm)
else
- match kind_of_term trm with
+ match kind trm with
| Meta _ ->
let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in
if !check && occur_meta sigma conclty then
@@ -372,7 +373,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
in
let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in
let sigma = check_conv_leq_goal env sigma trm conclty' conclty in
- let ans = if applicand == f && args == l then trm else Term.mkApp (applicand, args) in
+ let ans = if applicand == f && args == l then trm else mkApp (applicand, args) in
(acc'',conclty',sigma, ans)
| Proj (p,c) ->
@@ -394,7 +395,7 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
let lf' = Array.rev_of_list rbranches in
let ans =
if p' == p && c' == c && Array.equal (==) lf' lf then trm
- else Term.mkCase (ci,p',c',lf')
+ else mkCase (ci,p',c',lf')
in
(acc'',conclty',sigma, ans)
@@ -413,7 +414,7 @@ and mk_hdgoals sigma goal goalacc trm =
let hyps = Goal.V82.hyps sigma goal in
let mk_goal hyps concl =
Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in
- match kind_of_term trm with
+ match kind trm with
| Cast (c,_, ty) when isMeta c ->
check_typability env sigma ty;
let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in
@@ -433,7 +434,7 @@ and mk_hdgoals sigma goal goalacc trm =
else mk_hdgoals sigma goal goalacc f
in
let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in
- let ans = if applicand == f && args == l then trm else Term.mkApp (applicand, args) in
+ let ans = if applicand == f && args == l then trm else mkApp (applicand, args) in
(acc'',conclty',sigma, ans)
| Case (ci,p,c,lf) ->
@@ -447,7 +448,7 @@ and mk_hdgoals sigma goal goalacc trm =
let lf' = Array.rev_of_list rbranches in
let ans =
if p' == p && c' == c && Array.equal (==) lf' lf then trm
- else Term.mkCase (ci,p',c',lf')
+ else mkCase (ci,p',c',lf')
in
(acc'',conclty',sigma, ans)
@@ -468,12 +469,12 @@ and mk_arggoals sigma goal goalacc funty allargs =
let foldmap (goalacc, funty, sigma) harg =
let t = whd_all (Goal.V82.env sigma goal) sigma (EConstr.of_constr funty) in
let t = EConstr.Unsafe.to_constr t in
- let rec collapse t = match kind_of_term t with
+ let rec collapse t = match kind t with
| LetIn (_, c1, _, b) -> collapse (subst1 c1 b)
| _ -> t
in
let t = collapse t in
- match kind_of_term t with
+ match kind t with
| Prod (_, c1, b) ->
let (acc, hargty, sigma, arg) = mk_refgoals sigma goal goalacc c1 harg in
(acc, subst1 harg b, sigma), arg
diff --git a/proofs/logic.mli b/proofs/logic.mli
index 9d0756b33..7df7fd66b 100644
--- a/proofs/logic.mli
+++ b/proofs/logic.mli
@@ -9,7 +9,7 @@
(** Legacy proof engine. Do not use in newly written code. *)
open Names
-open Term
+open Constr
open Evd
open Proof_type
diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli
index 6e4ecd13b..ec176e547 100644
--- a/proofs/pfedit.mli
+++ b/proofs/pfedit.mli
@@ -10,7 +10,7 @@
open Loc
open Names
-open Term
+open Constr
open Environ
open Decl_kinds
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index 621178982..97faa1684 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -320,7 +320,7 @@ let constrain_variables init uctx =
let levels = Univ.Instance.levels (Univ.UContext.instance init) in
UState.constrain_variables levels uctx
-type closed_proof_output = (Term.constr * Safe_typing.private_constants) list * Evd.evar_universe_context
+type closed_proof_output = (Constr.t * Safe_typing.private_constants) list * Evd.evar_universe_context
let close_proof ~keep_body_ucst_separate ?feedback_id ~now
(fpl : closed_proof_output Future.computation) =
diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli
index 8c0f6ad85..6309f681f 100644
--- a/proofs/proof_global.mli
+++ b/proofs/proof_global.mli
@@ -86,7 +86,7 @@ val close_proof : keep_body_ucst_separate:bool -> Future.fix_exn -> closed_proof
* Both access the current proof state. The former is supposed to be
* chained with a computation that completed the proof *)
-type closed_proof_output = (Term.constr * Safe_typing.private_constants) list * Evd.evar_universe_context
+type closed_proof_output = (Constr.t * Safe_typing.private_constants) list * Evd.evar_universe_context
(* If allow_partial is set (default no) then an incomplete proof
* is allowed (no error), and a warn is given if the proof is complete. *)
diff --git a/proofs/proof_type.ml b/proofs/proof_type.ml
index 2ad5f607f..20293cb9b 100644
--- a/proofs/proof_type.ml
+++ b/proofs/proof_type.ml
@@ -9,7 +9,7 @@
(** Legacy proof engine. Do not use in newly written code. *)
open Evd
-open Term
+open Constr
(** This module defines the structure of proof tree and the tactic type. So, it
is used by [Proof_tree] and [Refiner] *)
diff --git a/proofs/redexpr.mli b/proofs/redexpr.mli
index ccc2440a2..43e598773 100644
--- a/proofs/redexpr.mli
+++ b/proofs/redexpr.mli
@@ -9,7 +9,7 @@
(** Interpretation layer of redexprs such as hnf, cbv, etc. *)
open Names
-open Term
+open Constr
open EConstr
open Pattern
open Genredexpr
diff --git a/proofs/refine.mli b/proofs/refine.mli
index 3b0a9e5b6..cfdcde36e 100644
--- a/proofs/refine.mli
+++ b/proofs/refine.mli
@@ -17,7 +17,7 @@ open Proofview
(** Printer used to print the constr which refine refines. *)
val pr_constr :
- (Environ.env -> Evd.evar_map -> Term.constr -> Pp.t) Hook.t
+ (Environ.env -> Evd.evar_map -> Constr.constr -> Pp.t) Hook.t
(** {7 Refinement primitives} *)
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index 3ff010fe3..9c8777c41 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -37,7 +37,7 @@ val tclIDTAC_MESSAGE : Pp.t -> tactic
val tclEVARS : evar_map -> tactic
val tclEVARUNIVCONTEXT : Evd.evar_universe_context -> tactic
-val tclPUSHCONTEXT : Evd.rigid -> Univ.universe_context_set -> tactic -> tactic
+val tclPUSHCONTEXT : Evd.rigid -> Univ.ContextSet.t -> tactic -> tactic
val tclPUSHEVARUNIVCONTEXT : Evd.evar_universe_context -> tactic
val tclPUSHCONSTRAINTS : Univ.constraints -> tactic
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index de9f8e700..6441cfd19 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open EConstr
open Evd
diff --git a/stm/asyncTaskQueue.ml b/stm/asyncTaskQueue.ml
index a356f32e9..e2bce1a96 100644
--- a/stm/asyncTaskQueue.ml
+++ b/stm/asyncTaskQueue.ml
@@ -59,7 +59,7 @@ module Make(T : Task) () = struct
type request = Request of T.request
type more_data =
- | MoreDataUnivLevel of Univ.universe_level list
+ | MoreDataUnivLevel of Univ.Level.t list
let slave_respond (Request r) =
let res = T.perform r in
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index ed612c0fc..e68087f14 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -10,6 +10,7 @@ open Equality
open Names
open Pp
open Term
+open Constr
open Termops
open CErrors
open Util
@@ -20,7 +21,7 @@ open Locus
type rew_rule = { rew_lemma: constr;
rew_type: types;
rew_pat: constr;
- rew_ctx: Univ.universe_context_set;
+ rew_ctx: Univ.ContextSet.t;
rew_l2r: bool;
rew_tac: Genarg.glob_generic_argument option }
diff --git a/tactics/autorewrite.mli b/tactics/autorewrite.mli
index edbb7c6b7..d2b5e070b 100644
--- a/tactics/autorewrite.mli
+++ b/tactics/autorewrite.mli
@@ -8,7 +8,7 @@
(** This files implements the autorewrite tactic. *)
-open Term
+open Constr
open Equality
(** Rewriting rules before tactic interpretation *)
@@ -28,7 +28,7 @@ val autorewrite_in : ?conds:conditions -> Names.Id.t -> unit Proofview.tactic ->
type rew_rule = { rew_lemma: constr;
rew_type: types;
rew_pat: constr;
- rew_ctx: Univ.universe_context_set;
+ rew_ctx: Univ.ContextSet.t;
rew_l2r: bool;
rew_tac: Genarg.glob_generic_argument option }
@@ -58,5 +58,5 @@ type hypinfo = {
val find_applied_relation :
?loc:Loc.t -> bool ->
- Environ.env -> Evd.evar_map -> Term.constr -> bool -> hypinfo
+ Environ.env -> Evd.evar_map -> constr -> bool -> hypinfo
diff --git a/tactics/elimschemes.ml b/tactics/elimschemes.ml
index 2d2a0c1b2..e427adb15 100644
--- a/tactics/elimschemes.ml
+++ b/tactics/elimschemes.ml
@@ -13,7 +13,8 @@
(* This file builds schemes related to case analysis and recursion schemes *)
-open Term
+open Sorts
+open Constr
open Indrec
open Declarations
open Typeops
diff --git a/tactics/elimschemes.mli b/tactics/elimschemes.mli
index e3fe7ddae..9c750e7ad 100644
--- a/tactics/elimschemes.mli
+++ b/tactics/elimschemes.mli
@@ -13,7 +13,7 @@ open Ind_tables
val optimize_non_type_induction_scheme :
'a Ind_tables.scheme_kind ->
Indrec.dep_flag ->
- Term.sorts_family ->
+ Sorts.family ->
'b ->
Names.inductive ->
(Constr.constr * Evd.evar_universe_context) * Safe_typing.private_constants
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index bfbac7787..d7667668e 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -48,6 +48,7 @@ open CErrors
open Util
open Names
open Term
+open Constr
open Vars
open Declarations
open Environ
@@ -106,8 +107,8 @@ let get_coq_eq ctx =
let univ_of_eq env eq =
let eq = EConstr.of_constr eq in
- match kind_of_term (EConstr.Unsafe.to_constr (Retyping.get_type_of env Evd.empty eq)) with
- | Prod (_,t,_) -> (match kind_of_term t with Sort (Type u) -> u | _ -> assert false)
+ match Constr.kind (EConstr.Unsafe.to_constr (Retyping.get_type_of env Evd.empty eq)) with
+ | Prod (_,t,_) -> (match Constr.kind t with Sort (Type u) -> u | _ -> assert false)
| _ -> assert false
(**********************************************************************)
@@ -141,7 +142,7 @@ let get_sym_eq_data env (ind,u) =
let paramsctxt = Vars.subst_instance_context u mib.mind_params_ctxt in
let paramsctxt1,_ =
List.chop (mib.mind_nparams-mip.mind_nrealargs) paramsctxt in
- if not (List.equal Term.eq_constr params2 constrargs) then
+ if not (List.equal Constr.equal params2 constrargs) then
error "Constructors arguments must repeat the parameters.";
(* nrealargs_ctxt and nrealargs are the same here *)
(specif,mip.mind_nrealargs,realsign,paramsctxt,paramsctxt1)
diff --git a/tactics/eqschemes.mli b/tactics/eqschemes.mli
index 4acfa7a28..90ae67c6c 100644
--- a/tactics/eqschemes.mli
+++ b/tactics/eqschemes.mli
@@ -9,7 +9,7 @@
(** This file builds schemes relative to equality inductive types *)
open Names
-open Term
+open Constr
open Environ
open Ind_tables
@@ -22,14 +22,14 @@ val rew_l2r_forward_dep_scheme_kind : individual scheme_kind
val rew_r2l_dep_scheme_kind : individual scheme_kind
val rew_r2l_scheme_kind : individual scheme_kind
-val build_r2l_rew_scheme : bool -> env -> inductive -> sorts_family ->
+val build_r2l_rew_scheme : bool -> env -> inductive -> Sorts.family ->
constr Evd.in_evar_universe_context
-val build_l2r_rew_scheme : bool -> env -> inductive -> sorts_family ->
+val build_l2r_rew_scheme : bool -> env -> inductive -> Sorts.family ->
constr Evd.in_evar_universe_context * Safe_typing.private_constants
val build_r2l_forward_rew_scheme :
- bool -> env -> inductive -> sorts_family -> constr Evd.in_evar_universe_context
+ bool -> env -> inductive -> Sorts.family -> constr Evd.in_evar_universe_context
val build_l2r_forward_rew_scheme :
- bool -> env -> inductive -> sorts_family -> constr Evd.in_evar_universe_context
+ bool -> env -> inductive -> Sorts.family -> constr Evd.in_evar_universe_context
(** Builds a symmetry scheme for a symmetrical equality type *)
@@ -43,5 +43,5 @@ val sym_involutive_scheme_kind : individual scheme_kind
(** Builds a congruence scheme for an equality type *)
val congr_scheme_kind : individual scheme_kind
-val build_congr : env -> constr * constr * Univ.universe_context_set -> inductive ->
+val build_congr : env -> constr * constr * Univ.ContextSet.t -> inductive ->
constr Evd.in_evar_universe_context
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 3ccbab874..c7c53b393 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -128,14 +128,14 @@ type hints_path = global_reference hints_path_gen
type hint_term =
| IsGlobRef of global_reference
- | IsConstr of constr * Univ.universe_context_set
+ | IsConstr of constr * Univ.ContextSet.t
type 'a with_uid = {
obj : 'a;
uid : KerName.t;
}
-type raw_hint = constr * types * Univ.universe_context_set
+type raw_hint = constr * types * Univ.ContextSet.t
type hint = (raw_hint * clausenv) hint_ast with_uid
diff --git a/tactics/hints.mli b/tactics/hints.mli
index 44e5370e9..22df29b80 100644
--- a/tactics/hints.mli
+++ b/tactics/hints.mli
@@ -42,7 +42,7 @@ type 'a hint_ast =
| Extern of Genarg.glob_generic_argument (* Hint Extern *)
type hint
-type raw_hint = constr * types * Univ.universe_context_set
+type raw_hint = constr * types * Univ.ContextSet.t
type 'a hints_path_atom_gen =
| PathHints of 'a list
@@ -146,7 +146,7 @@ type hint_info = (patvar list * constr_pattern) hint_info_gen
type hint_term =
| IsGlobRef of global_reference
- | IsConstr of constr * Univ.universe_context_set
+ | IsConstr of constr * Univ.ContextSet.t
type hints_entry =
| HintsResolveEntry of
@@ -193,7 +193,7 @@ val prepare_hint : bool (* Check no remaining evars *) ->
*)
val make_exact_entry : env -> evar_map -> hint_info -> polymorphic -> ?name:hints_path_atom ->
- (constr * types * Univ.universe_context_set) -> hint_entry
+ (constr * types * Univ.ContextSet.t) -> hint_entry
(** [make_apply_entry (eapply,hnf,verbose) info (c,cty,ctx))].
[eapply] is true if this hint will be used only with EApply;
@@ -211,7 +211,7 @@ val make_exact_entry : env -> evar_map -> hint_info -> polymorphic -> ?name:hint
val make_apply_entry :
env -> evar_map -> bool * bool * bool -> hint_info -> polymorphic -> ?name:hints_path_atom ->
- (constr * types * Univ.universe_context_set) -> hint_entry
+ (constr * types * Univ.ContextSet.t) -> hint_entry
(** A constr which is Hint'ed will be:
- (1) used as an Exact, if it does not start with a product
diff --git a/tactics/ind_tables.ml b/tactics/ind_tables.ml
index 240b5a7e0..e7fa555c2 100644
--- a/tactics/ind_tables.ml
+++ b/tactics/ind_tables.ml
@@ -17,7 +17,7 @@ open Mod_subst
open Libobject
open Nameops
open Declarations
-open Term
+open Constr
open CErrors
open Util
open Declare
diff --git a/tactics/ind_tables.mli b/tactics/ind_tables.mli
index 232a193ac..d73595a2f 100644
--- a/tactics/ind_tables.mli
+++ b/tactics/ind_tables.mli
@@ -6,8 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
open Names
+open Constr
open Declare
(** This module provides support for registering inductive scheme builders,
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index fe733899f..cea6ccc30 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -10,7 +10,7 @@ open Pp
open CErrors
open Util
open Names
-open Term
+open Constr
open EConstr
open Termops
open Declarations
@@ -224,9 +224,8 @@ let compute_induction_names = compute_induction_names_gen true
(* Compute the let-in signature of case analysis or standard induction scheme *)
let compute_constructor_signatures isrec ((_,k as ity),u) =
- let open Term in
let rec analrec c recargs =
- match kind_of_term c, recargs with
+ match Constr.kind c, recargs with
| Prod (_,_,c), recarg::rest ->
let rest = analrec c rest in
begin match Declareops.dest_recarg recarg with
@@ -242,7 +241,7 @@ let compute_constructor_signatures isrec ((_,k as ity),u) =
let (mib,mip) = Global.lookup_inductive ity in
let n = mib.mind_nparams in
let lc =
- Array.map (fun c -> snd (decompose_prod_n_assum n c)) mip.mind_nf_lc in
+ Array.map (fun c -> snd (Term.decompose_prod_n_assum n c)) mip.mind_nf_lc in
let lrecargs = Declareops.dest_subterms mip.mind_recargs in
Array.map2 analrec lc lrecargs
@@ -472,7 +471,7 @@ module New = struct
let evi = Evd.find sigma evk in
match Evd.evar_body evi with
| Evd.Evar_empty -> Some (evk,evi)
- | Evd.Evar_defined c -> match Term.kind_of_term c with
+ | Evd.Evar_defined c -> match Constr.kind c with
| Term.Evar (evk,l) -> is_undefined_up_to_restriction sigma evk
| _ ->
(* We make the assumption that there is no way to refine an
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 3abd42d46..f5c209c74 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open EConstr
open Tacmach
open Proof_type
@@ -127,9 +127,9 @@ val compute_constructor_signatures : rec_flag -> inductive * 'a -> bool list arr
val compute_induction_names :
bool list array -> or_and_intro_pattern option -> intro_patterns array
-val elimination_sort_of_goal : goal sigma -> sorts_family
-val elimination_sort_of_hyp : Id.t -> goal sigma -> sorts_family
-val elimination_sort_of_clause : Id.t option -> goal sigma -> sorts_family
+val elimination_sort_of_goal : goal sigma -> Sorts.family
+val elimination_sort_of_hyp : Id.t -> goal sigma -> Sorts.family
+val elimination_sort_of_clause : Id.t option -> goal sigma -> Sorts.family
val pf_with_evars : (goal sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tactic
val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> tactic
@@ -243,9 +243,9 @@ module New : sig
val tryAllHypsAndConcl : (Id.t option -> unit tactic) -> unit tactic
val onClause : (Id.t option -> unit tactic) -> clause -> unit tactic
- val elimination_sort_of_goal : 'a Proofview.Goal.t -> sorts_family
- val elimination_sort_of_hyp : Id.t -> 'a Proofview.Goal.t -> sorts_family
- val elimination_sort_of_clause : Id.t option -> 'a Proofview.Goal.t -> sorts_family
+ val elimination_sort_of_goal : 'a Proofview.Goal.t -> Sorts.family
+ val elimination_sort_of_hyp : Id.t -> 'a Proofview.Goal.t -> Sorts.family
+ val elimination_sort_of_clause : Id.t option -> 'a Proofview.Goal.t -> Sorts.family
val elimination_then :
(branch_args -> unit Proofview.tactic) ->
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 866daa904..ba244a794 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -14,6 +14,7 @@ open Util
open Names
open Nameops
open Term
+open Constr
open Termops
open Environ
open EConstr
@@ -1290,7 +1291,7 @@ let check_unresolved_evars_of_metas sigma clenv =
(* Refiner.pose_all_metas_as_evars are resolved *)
List.iter (fun (mv,b) -> match b with
| Clval (_,(c,_),_) ->
- (match kind_of_term c.rebus with
+ (match Constr.kind c.rebus with
| Evar (evk,_) when Evd.is_undefined clenv.evd evk
&& not (Evd.mem sigma evk) ->
error_uninstantiated_metas (mkMeta mv) clenv
@@ -4958,7 +4959,7 @@ let interpretable_as_section_decl evd d1 d2 =
let rec decompose len c t accu =
let open Context.Rel.Declaration in
if len = 0 then (c, t, accu)
- else match kind_of_term c, kind_of_term t with
+ else match Constr.kind c, Constr.kind t with
| Lambda (na, u, c), Prod (_, _, t) ->
decompose (pred len) c t (LocalAssum (na, u) :: accu)
| LetIn (na, b, u, c), LetIn (_, _, _, t) ->
@@ -4966,7 +4967,7 @@ let rec decompose len c t accu =
| _ -> assert false
let rec shrink ctx sign c t accu =
- let open Term in
+ let open Constr in
let open CVars in
match ctx, sign with
| [], [] -> (c, t, accu)
@@ -4976,8 +4977,8 @@ let rec shrink ctx sign c t accu =
let t = subst1 mkProp t in
shrink ctx sign c t accu
else
- let c = mkLambda_or_LetIn p c in
- let t = mkProd_or_LetIn p t in
+ let c = Term.mkLambda_or_LetIn p c in
+ let t = Term.mkProd_or_LetIn p t in
let accu = if RelDecl.is_local_assum p
then mkVar (NamedDecl.get_id decl) :: accu
else accu
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 98cf1b437..83fc655f1 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -8,7 +8,7 @@
open Loc
open Names
-open Term
+open Constr
open EConstr
open Environ
open Proof_type
diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml
index 64ba38a51..6c8130d79 100644
--- a/tactics/term_dnet.ml
+++ b/tactics/term_dnet.ml
@@ -8,7 +8,7 @@
(*i*)
open Util
-open Term
+open Constr
open Names
open Globnames
open Mod_subst
@@ -257,7 +257,7 @@ struct
let pat_of_constr c : term_pattern =
(** To each evar we associate a unique identifier. *)
let metas = ref Evar.Map.empty in
- let rec pat_of_constr c = match kind_of_term c with
+ let rec pat_of_constr c = match Constr.kind c with
| Rel _ -> Term DRel
| Sort _ -> Term DSort
| Var i -> Term (DRef (VarRef i))
@@ -290,7 +290,7 @@ struct
| Proj (p,c) ->
Term (DApp (Term (DRef (ConstRef (Projection.constant p))), pat_of_constr c))
- and ctx_of_constr ctx c = match kind_of_term c with
+ and ctx_of_constr ctx c = match Constr.kind c with
| Prod (_,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t,None),ctx))) c
| LetIn(_,d,t,c) -> ctx_of_constr (Term(DCons((pat_of_constr t, Some (pat_of_constr d)),ctx))) c
| _ -> ctx,pat_of_constr c
diff --git a/tactics/term_dnet.mli b/tactics/term_dnet.mli
index 16122fa5e..db7da18ba 100644
--- a/tactics/term_dnet.mli
+++ b/tactics/term_dnet.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open Constr
open Mod_subst
(** Dnets on constr terms.
diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml
index 09e645eea..d22024568 100644
--- a/vernac/assumptions.ml
+++ b/vernac/assumptions.ml
@@ -18,7 +18,7 @@ open Pp
open CErrors
open Util
open Names
-open Term
+open Constr
open Declarations
open Mod_subst
open Globnames
@@ -163,7 +163,7 @@ let label_of = function
let fold_constr_with_full_binders g f n acc c =
let open Context.Rel.Declaration in
- match kind_of_term c with
+ match Constr.kind c with
| Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ -> acc
| Cast (c,_, t) -> f n (f n acc c) t
| Prod (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
@@ -182,7 +182,7 @@ let fold_constr_with_full_binders g f n acc c =
let fd = Array.map2 (fun t b -> (t,b)) tl bl in
Array.fold_left (fun acc (t,b) -> f n' (f n acc t) b) acc fd
-let rec traverse current ctx accu t = match kind_of_term t with
+let rec traverse current ctx accu t = match Constr.kind t with
| Var id ->
let body () = id |> Global.lookup_named |> NamedDecl.get_value in
traverse_object accu body (VarRef id)
diff --git a/vernac/assumptions.mli b/vernac/assumptions.mli
index 77eb968d4..afe932ead 100644
--- a/vernac/assumptions.mli
+++ b/vernac/assumptions.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Globnames
open Printer
@@ -28,4 +28,4 @@ val traverse :
{!traverse} also applies. *)
val assumptions :
?add_opaque:bool -> ?add_transparent:bool -> transparent_state ->
- global_reference -> constr -> Term.types ContextObjectMap.t
+ global_reference -> constr -> types ContextObjectMap.t
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 539e5550f..3cf181441 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Pp
open Term
+open Constr
open Vars
open Termops
open Declarations
@@ -532,7 +533,7 @@ let eqI ind l =
try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff
with Not_found -> user_err ~hdr:"AutoIndDecl.eqI"
(str "The boolean equality on " ++ MutInd.print (fst ind) ++ str " is needed.");
- in (if Array.equal Term.eq_constr eA [||] then e else mkApp(e,eA)), eff
+ in (if Array.equal Constr.equal eA [||] then e else mkApp(e,eA)), eff
(**********************************************************************)
(* Boolean->Leibniz *)
diff --git a/vernac/class.ml b/vernac/class.ml
index 061f3efcc..f26599973 100644
--- a/vernac/class.ml
+++ b/vernac/class.ml
@@ -11,6 +11,7 @@ open Util
open Pp
open Names
open Term
+open Constr
open Vars
open Termops
open Entries
@@ -148,7 +149,7 @@ let get_target t ind =
let prods_of t =
- let rec aux acc d = match kind_of_term d with
+ let rec aux acc d = match Constr.kind d with
| Prod (_,c1,c2) -> aux (c1::acc) c2
| Cast (c,_,_) -> aux acc c
| _ -> (d,acc)
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 9a8fc9bc2..22117f7e1 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -9,6 +9,7 @@
(*i*)
open Names
open Term
+open Constr
open Vars
open Environ
open Nametab
diff --git a/vernac/classes.mli b/vernac/classes.mli
index fcdb5c3bc..c0f03227c 100644
--- a/vernac/classes.mli
+++ b/vernac/classes.mli
@@ -34,7 +34,7 @@ val declare_instance_constant :
bool -> (* polymorphic *)
Evd.evar_map -> (* Universes *)
Constr.t -> (** body *)
- Term.types -> (** type *)
+ Constr.types -> (** type *)
Names.Id.t
val new_instance :
diff --git a/vernac/command.ml b/vernac/command.ml
index f58ed065c..db3fa1955 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -10,6 +10,7 @@ open Pp
open CErrors
open Util
open Term
+open Constr
open Vars
open Termops
open Environ
@@ -44,7 +45,7 @@ let do_constraint poly l = Declare.do_constraint poly l
let rec under_binders env sigma f n c =
if Int.equal n 0 then f env sigma (EConstr.of_constr c) else
- match kind_of_term c with
+ match Constr.kind c with
| Lambda (x,t,c) ->
mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
| LetIn (x,b,t,c) ->
@@ -652,7 +653,7 @@ let extract_mutual_inductive_declaration_components indl =
let is_recursive mie =
let rec is_recursive_constructor lift typ =
- match Term.kind_of_term typ with
+ match Constr.kind typ with
| Prod (_,arg,rest) ->
not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) ||
is_recursive_constructor (lift+1) rest
diff --git a/vernac/command.mli b/vernac/command.mli
index 26b1d1aaf..5415d3308 100644
--- a/vernac/command.mli
+++ b/vernac/command.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Entries
open Libnames
open Globnames
diff --git a/vernac/declareDef.mli b/vernac/declareDef.mli
index 5dea0ba27..01a87818a 100644
--- a/vernac/declareDef.mli
+++ b/vernac/declareDef.mli
@@ -15,5 +15,5 @@ val declare_definition : Id.t -> definition_kind ->
Safe_typing.private_constants Entries.definition_entry -> Universes.universe_binders -> Impargs.manual_implicits ->
Globnames.global_reference Lemmas.declaration_hook -> Globnames.global_reference
-val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t ->
+val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.UContext.t -> Id.t ->
Safe_typing.private_constants Entries.proof_output -> Constr.types -> Impargs.manual_implicits -> Globnames.global_reference
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index 4bdc93a36..c0ddc7e2c 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -21,6 +21,7 @@ open Names
open Declarations
open Entries
open Term
+open Constr
open Inductive
open Decl_kinds
open Indrec
@@ -458,7 +459,7 @@ let build_combined_scheme env schemes =
let find_inductive ty =
let (ctx, arity) = decompose_prod ty in
let (_, last) = List.hd ctx in
- match kind_of_term last with
+ match Constr.kind last with
| App (ind, args) ->
let ind = destInd ind in
let (_,spec) = Inductive.lookup_mind_specif env (fst ind) in
diff --git a/vernac/indschemes.mli b/vernac/indschemes.mli
index 659f12936..4b31389ab 100644
--- a/vernac/indschemes.mli
+++ b/vernac/indschemes.mli
@@ -8,7 +8,7 @@
open Loc
open Names
-open Term
+open Constr
open Environ
open Vernacexpr
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 22f0d199c..be9de5b30 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -14,6 +14,7 @@ open Util
open Pp
open Names
open Term
+open Constr
open Declarations
open Declareops
open Entries
@@ -62,7 +63,7 @@ let adjust_guardness_conditions const = function
{ const with const_entry_body =
Future.chain const.const_entry_body
(fun ((body, ctx), eff) ->
- match kind_of_term body with
+ match Constr.kind body with
| Fix ((nv,0),(_,_,fixdefs as fixdecls)) ->
(* let possible_indexes =
List.map2 (fun i c -> match i with Some i -> i | None ->
@@ -97,7 +98,7 @@ let find_mutually_recursive_statements thms =
let ind_hyps =
List.flatten (List.map_i (fun i decl ->
let t = RelDecl.get_type decl in
- match kind_of_term t with
+ match Constr.kind t with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
mind.mind_finite <> Decl_kinds.CoFinite ->
@@ -107,7 +108,7 @@ let find_mutually_recursive_statements thms =
let ind_ccl =
let cclenv = push_rel_context hyps (Global.env()) in
let whnf_ccl,_ = whd_all_stack cclenv Evd.empty (EConstr.of_constr ccl) in
- match kind_of_term (EConstr.Unsafe.to_constr whnf_ccl) with
+ match Constr.kind (EConstr.Unsafe.to_constr whnf_ccl) with
| Ind ((kn,_ as ind),u) when
let mind = Global.lookup_mind kn in
Int.equal mind.mind_ntypes n && mind.mind_finite == Decl_kinds.CoFinite ->
@@ -246,7 +247,7 @@ let save_remaining_recthms (locality,p,kind) norm ctx binders body opaq i (id,(t
| Some body ->
let body = norm body in
let k = Kindops.logical_kind_of_goal_kind kind in
- let rec body_i t = match kind_of_term t with
+ let rec body_i t = match Constr.kind t with
| Fix ((nv,0),decls) -> mkFix ((nv,i),decls)
| CoFix (0,decls) -> mkCoFix (i,decls)
| LetIn(na,t1,ty,t2) -> mkLetIn (na,t1,ty, body_i t2)
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 1e23c7314..1f46a385d 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Decl_kinds
type 'a declaration_hook
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 785c842ba..e23146273 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -13,6 +13,7 @@ open Declare
*)
open Term
+open Constr
open Vars
open Names
open Evd
@@ -55,7 +56,7 @@ let subst_evar_constr evs n idf t =
let seen = ref Int.Set.empty in
let transparent = ref Id.Set.empty in
let evar_info id = List.assoc_f Evar.equal id evs in
- let rec substrec (depth, fixrels) c = match kind_of_term c with
+ let rec substrec (depth, fixrels) c = match Constr.kind c with
| Evar (k, args) ->
let { ev_name = (id, idstr) ;
ev_hyps = hyps ; ev_chop = chop } =
@@ -85,15 +86,15 @@ let subst_evar_constr evs n idf t =
in aux hyps args []
in
if List.exists
- (fun x -> match kind_of_term x with
+ (fun x -> match Constr.kind x with
| Rel n -> Int.List.mem n fixrels
| _ -> false) args
then
transparent := Id.Set.add idstr !transparent;
mkApp (idf idstr, Array.of_list args)
| Fix _ ->
- map_constr_with_binders succfix substrec (depth, 1 :: fixrels) c
- | _ -> map_constr_with_binders succfix substrec (depth, fixrels) c
+ Constr.map_with_binders succfix substrec (depth, 1 :: fixrels) c
+ | _ -> Constr.map_with_binders succfix substrec (depth, fixrels) c
in
let t' = substrec (0, []) t in
t', !seen, !transparent
@@ -103,9 +104,9 @@ let subst_evar_constr evs n idf t =
where n binders were passed through. *)
let subst_vars acc n t =
let var_index id = Util.List.index Id.equal id acc in
- let rec substrec depth c = match kind_of_term c with
+ let rec substrec depth c = match Constr.kind c with
| Var v -> (try mkRel (depth + (var_index v)) with Not_found -> c)
- | _ -> map_constr_with_binders succ substrec depth c
+ | _ -> Constr.map_with_binders succ substrec depth c
in
substrec 0 t
@@ -144,7 +145,7 @@ let rec chop_product n t =
let pop t = Vars.lift (-1) t in
if Int.equal n 0 then Some t
else
- match kind_of_term t with
+ match Constr.kind t with
| Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (pop b) else None
| _ -> None
@@ -273,7 +274,7 @@ let explain_no_obligations = function
| None -> str "No obligations remaining"
type obligation_info =
- (Names.Id.t * Term.types * Evar_kinds.t Loc.located *
+ (Names.Id.t * types * Evar_kinds.t Loc.located *
(bool * Evar_kinds.obligation_definition_status)
* Int.Set.t * unit Proofview.tactic option) array
@@ -384,7 +385,7 @@ let subst_deps expand obls deps t =
(Vars.replace_vars (List.map (fun (n, (_, b)) -> n, b) osubst) t)
let rec prod_app t n =
- match kind_of_term (EConstr.Unsafe.to_constr (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t))) (** FIXME *) with
+ match Constr.kind (EConstr.Unsafe.to_constr (Termops.strip_outer_cast Evd.empty (EConstr.of_constr t))) (** FIXME *) with
| Prod (_,_,b) -> subst1 n b
| LetIn (_, b, t, b') -> prod_app (subst1 b b') n
| _ ->
@@ -400,13 +401,13 @@ let replace_appvars subst =
let f, l = decompose_app c in
if isVar f then
try
- let c' = List.map (map_constr aux) l in
+ let c' = List.map (Constr.map aux) l in
let (t, b) = Id.List.assoc (destVar f) subst in
mkApp (delayed_force hide_obligation,
[| prod_applist t c'; applistc b c' |])
- with Not_found -> map_constr aux c
- else map_constr aux c
- in map_constr aux
+ with Not_found -> Constr.map aux c
+ else Constr.map aux c
+ in Constr.map aux
let subst_prog expand obls ints prg =
let subst = obl_substitution expand obls ints in
@@ -490,7 +491,7 @@ let declare_definition prg =
cst
let rec lam_index n t acc =
- match kind_of_term t with
+ match Constr.kind t with
| Lambda (Name n', _, _) when Id.equal n n' ->
acc
| Lambda (_, _, b) ->
@@ -566,9 +567,9 @@ let declare_mutual_definition l =
let decompose_lam_prod c ty =
let open Context.Rel.Declaration in
let rec aux ctx c ty =
- match kind_of_term c, kind_of_term ty with
+ match Constr.kind c, Constr.kind ty with
| LetIn (x, b, t, c), LetIn (x', b', t', ty)
- when eq_constr b b' && eq_constr t t' ->
+ when Constr.equal b b' && Constr.equal t t' ->
let ctx' = Context.Rel.add (LocalDef (x,b',t')) ctx in
aux ctx' c ty
| _, LetIn (x', b', t', ty) ->
diff --git a/vernac/obligations.mli b/vernac/obligations.mli
index 11c2553ae..d037fdcd8 100644
--- a/vernac/obligations.mli
+++ b/vernac/obligations.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Environ
-open Term
+open Constr
open Evd
open Names
open Globnames
@@ -39,7 +39,7 @@ val eterm_obligations : env -> Id.t -> evar_map -> int ->
translation from obligation identifiers to constrs, new term, new type *)
type obligation_info =
- (Id.t * Term.types * Evar_kinds.t Loc.located *
+ (Id.t * types * Evar_kinds.t Loc.located *
(bool * Evar_kinds.obligation_definition_status) * Int.Set.t * unit Proofview.tactic option) array
(* ident, type, location, (opaque or transparent, expand or define),
dependencies, tactic to solve it *)
@@ -51,13 +51,13 @@ type progress = (* Resolution status of a program *)
val default_tactic : unit Proofview.tactic ref
-val add_definition : Names.Id.t -> ?term:Term.constr -> Term.types ->
+val add_definition : Names.Id.t -> ?term:constr -> types ->
Evd.evar_universe_context ->
?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *)
?implicits:(Constrexpr.explicitation * (bool * bool * bool)) list ->
?kind:Decl_kinds.definition_kind ->
?tactic:unit Proofview.tactic ->
- ?reduce:(Term.constr -> Term.constr) ->
+ ?reduce:(constr -> constr) ->
?hook:(Evd.evar_universe_context -> unit) Lemmas.declaration_hook -> ?opaque:bool -> obligation_info -> progress
type notations =
@@ -68,13 +68,13 @@ type fixpoint_kind =
| IsCoFixpoint
val add_mutual_definitions :
- (Names.Id.t * Term.constr * Term.types *
+ (Names.Id.t * constr * types *
(Constrexpr.explicitation * (bool * bool * bool)) list * obligation_info) list ->
Evd.evar_universe_context ->
?univdecl:Univdecls.universe_decl -> (* Universe binders and constraints *)
?tactic:unit Proofview.tactic ->
?kind:Decl_kinds.definition_kind ->
- ?reduce:(Term.constr -> Term.constr) ->
+ ?reduce:(constr -> constr) ->
?hook:(Evd.evar_universe_context -> unit) Lemmas.declaration_hook -> ?opaque:bool ->
notations ->
fixpoint_kind -> unit
diff --git a/vernac/record.ml b/vernac/record.ml
index 5533fe5b3..1fd43624a 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -13,6 +13,7 @@ open Names
open Globnames
open Nameops
open Term
+open Constr
open Vars
open Environ
open Declarations
@@ -229,7 +230,7 @@ exception NotDefinable of record_error
let subst_projection fid l c =
let lv = List.length l in
let bad_projs = ref [] in
- let rec substrec depth c = match kind_of_term c with
+ let rec substrec depth c = match Constr.kind c with
| Rel k ->
(* We are in context [[params;fields;x:ind;...depth...]] *)
if k <= depth+1 then
@@ -244,7 +245,7 @@ let subst_projection fid l c =
" field which has no name.")
else
mkRel (k-lv)
- | _ -> map_constr_with_binders succ substrec depth c
+ | _ -> Constr.map_with_binders succ substrec depth c
in
let c' = lift 1 c in (* to get [c] defined in ctxt [[params;fields;x:ind]] *)
let c'' = substrec 0 c' in
diff --git a/vernac/record.mli b/vernac/record.mli
index 1bcbf39b7..33c2fba89 100644
--- a/vernac/record.mli
+++ b/vernac/record.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Vernacexpr
open Constrexpr
open Impargs
diff --git a/vernac/search.ml b/vernac/search.ml
index 0f56f81e7..6da6a0c2d 100644
--- a/vernac/search.ml
+++ b/vernac/search.ml
@@ -9,7 +9,7 @@
open Pp
open Util
open Names
-open Term
+open Constr
open Declarations
open Libobject
open Environ
diff --git a/vernac/search.mli b/vernac/search.mli
index db54d732b..2eda3980a 100644
--- a/vernac/search.mli
+++ b/vernac/search.mli
@@ -7,7 +7,7 @@
(************************************************************************)
open Names
-open Term
+open Constr
open Environ
open Pattern
open Globnames