aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
authorGravatar Maxime Dénès <mail@maximedenes.fr>2017-04-11 00:28:47 +0200
committerGravatar Maxime Dénès <mail@maximedenes.fr>2017-04-11 00:28:47 +0200
commit835be3a05e28eb6e26f703a034f22b2c6c61acaa (patch)
tree00ecf04840ba027c3c71f8503d9811c8a5dc1d2e
parent0980dbb1740c8d48d8ff0c516929f27f8cea854d (diff)
parent2e6a89238dc7197057d0da80a16f4b4b1e41bfd8 (diff)
Merge PR#379: Introducing evar-insensitive constrs
-rw-r--r--CHANGES11
-rw-r--r--dev/db1
-rw-r--r--dev/top_printers.ml25
-rw-r--r--engine/eConstr.ml782
-rw-r--r--engine/eConstr.mli282
-rw-r--r--engine/engine.mllib3
-rw-r--r--engine/evarutil.ml204
-rw-r--r--engine/evarutil.mli38
-rw-r--r--engine/evd.ml352
-rw-r--r--engine/evd.mli35
-rw-r--r--engine/namegen.ml118
-rw-r--r--engine/namegen.mli44
-rw-r--r--engine/proofview.ml28
-rw-r--r--engine/proofview.mli13
-rw-r--r--engine/sigma.ml4
-rw-r--r--engine/sigma.mli2
-rw-r--r--engine/termops.ml763
-rw-r--r--engine/termops.mli220
-rw-r--r--engine/universes.ml9
-rw-r--r--engine/universes.mli6
-rw-r--r--ide/ide_slave.ml2
-rw-r--r--interp/constrextern.ml8
-rw-r--r--interp/constrintern.ml7
-rw-r--r--interp/constrintern.mli24
-rw-r--r--interp/notation.ml4
-rw-r--r--interp/notation_ops.ml1
-rw-r--r--interp/reserve.ml2
-rw-r--r--interp/stdarg.mli1
-rw-r--r--intf/pattern.mli4
-rw-r--r--intf/tactypes.mli4
-rw-r--r--kernel/constr.ml18
-rw-r--r--kernel/constr.mli26
-rw-r--r--kernel/context.ml69
-rw-r--r--kernel/context.mli160
-rw-r--r--kernel/environ.ml16
-rw-r--r--kernel/environ.mli19
-rw-r--r--kernel/indtypes.ml2
-rw-r--r--kernel/inductive.ml4
-rw-r--r--kernel/term.ml13
-rw-r--r--kernel/term.mli14
-rw-r--r--kernel/type_errors.ml58
-rw-r--r--kernel/type_errors.mli58
-rw-r--r--kernel/vars.mli2
-rw-r--r--library/impargs.ml8
-rw-r--r--library/keys.ml4
-rw-r--r--library/keys.mli2
-rw-r--r--plugins/btauto/refl_btauto.ml26
-rw-r--r--plugins/cc/ccalgo.ml20
-rw-r--r--plugins/cc/cctac.ml138
-rw-r--r--plugins/cc/cctac.mli2
-rw-r--r--plugins/derive/derive.ml4
-rw-r--r--plugins/extraction/extraction.ml43
-rw-r--r--plugins/firstorder/formula.ml27
-rw-r--r--plugins/firstorder/instances.ml15
-rw-r--r--plugins/firstorder/rules.ml25
-rw-r--r--plugins/firstorder/sequent.ml6
-rw-r--r--plugins/firstorder/unify.ml18
-rw-r--r--plugins/fourier/fourierR.ml39
-rw-r--r--plugins/funind/functional_principles_proofs.ml277
-rw-r--r--plugins/funind/functional_principles_proofs.mli6
-rw-r--r--plugins/funind/functional_principles_types.ml64
-rw-r--r--plugins/funind/functional_principles_types.mli2
-rw-r--r--plugins/funind/g_indfun.ml411
-rw-r--r--plugins/funind/glob_term_to_relation.ml36
-rw-r--r--plugins/funind/indfun.ml55
-rw-r--r--plugins/funind/indfun.mli4
-rw-r--r--plugins/funind/indfun_common.ml55
-rw-r--r--plugins/funind/indfun_common.mli23
-rw-r--r--plugins/funind/invfun.ml148
-rw-r--r--plugins/funind/merge.ml19
-rw-r--r--plugins/funind/recdef.ml281
-rw-r--r--plugins/funind/recdef.mli2
-rw-r--r--plugins/ltac/evar_tactics.ml38
-rw-r--r--plugins/ltac/evar_tactics.mli4
-rw-r--r--plugins/ltac/extraargs.mli4
-rw-r--r--plugins/ltac/extratactics.ml4166
-rw-r--r--plugins/ltac/g_auto.ml47
-rw-r--r--plugins/ltac/g_class.ml416
-rw-r--r--plugins/ltac/pptactic.ml37
-rw-r--r--plugins/ltac/pptactic.mli6
-rw-r--r--plugins/ltac/rewrite.ml238
-rw-r--r--plugins/ltac/rewrite.mli7
-rw-r--r--plugins/ltac/taccoerce.ml58
-rw-r--r--plugins/ltac/taccoerce.mli25
-rw-r--r--plugins/ltac/tacexpr.mli6
-rw-r--r--plugins/ltac/tacinterp.ml132
-rw-r--r--plugins/ltac/tacinterp.mli3
-rw-r--r--plugins/ltac/tactic_debug.ml14
-rw-r--r--plugins/ltac/tactic_debug.mli7
-rw-r--r--plugins/ltac/tactic_matching.ml8
-rw-r--r--plugins/ltac/tactic_matching.mli10
-rw-r--r--plugins/ltac/tauto.ml37
-rw-r--r--plugins/micromega/coq_micromega.ml268
-rw-r--r--plugins/nsatz/g_nsatz.ml42
-rw-r--r--plugins/nsatz/nsatz.ml1
-rw-r--r--plugins/omega/coq_omega.ml215
-rw-r--r--plugins/quote/g_quote.ml42
-rw-r--r--plugins/quote/quote.ml115
-rw-r--r--plugins/romega/const_omega.ml2
-rw-r--r--plugins/romega/refl_omega.ml11
-rw-r--r--plugins/rtauto/refl_tauto.ml26
-rw-r--r--plugins/rtauto/refl_tauto.mli6
-rw-r--r--plugins/setoid_ring/g_newring.ml410
-rw-r--r--plugins/setoid_ring/newring.ml135
-rw-r--r--plugins/setoid_ring/newring.mli43
-rw-r--r--plugins/ssrmatching/ssrmatching.ml4105
-rw-r--r--plugins/ssrmatching/ssrmatching.mli6
-rw-r--r--pretyping/cases.ml349
-rw-r--r--pretyping/cases.mli22
-rw-r--r--pretyping/cbv.ml3
-rw-r--r--pretyping/cbv.mli4
-rw-r--r--pretyping/classops.ml25
-rw-r--r--pretyping/classops.mli7
-rw-r--r--pretyping/coercion.ml83
-rw-r--r--pretyping/coercion.mli1
-rw-r--r--pretyping/constr_matching.ml105
-rw-r--r--pretyping/constr_matching.mli3
-rw-r--r--pretyping/detyping.ml127
-rw-r--r--pretyping/detyping.mli7
-rw-r--r--pretyping/evarconv.ml220
-rw-r--r--pretyping/evarconv.mli9
-rw-r--r--pretyping/evardefine.ml70
-rw-r--r--pretyping/evardefine.mli3
-rw-r--r--pretyping/evarsolve.ml662
-rw-r--r--pretyping/evarsolve.mli15
-rw-r--r--pretyping/find_subterm.ml34
-rw-r--r--pretyping/find_subterm.mli8
-rw-r--r--pretyping/indrec.ml75
-rw-r--r--pretyping/inductiveops.ml128
-rw-r--r--pretyping/inductiveops.mli40
-rw-r--r--pretyping/miscops.ml13
-rw-r--r--pretyping/miscops.mli5
-rw-r--r--pretyping/nativenorm.ml103
-rw-r--r--pretyping/nativenorm.mli3
-rw-r--r--pretyping/patternops.ml12
-rw-r--r--pretyping/patternops.mli5
-rw-r--r--pretyping/pretype_errors.ml19
-rw-r--r--pretyping/pretype_errors.mli18
-rw-r--r--pretyping/pretyping.ml208
-rw-r--r--pretyping/pretyping.mli5
-rw-r--r--pretyping/program.ml7
-rw-r--r--pretyping/program.mli2
-rw-r--r--pretyping/recordops.ml23
-rw-r--r--pretyping/recordops.mli2
-rw-r--r--pretyping/reductionops.ml410
-rw-r--r--pretyping/reductionops.mli52
-rw-r--r--pretyping/retyping.ml102
-rw-r--r--pretyping/retyping.mli3
-rw-r--r--pretyping/tacred.ml330
-rw-r--r--pretyping/tacred.mli7
-rw-r--r--pretyping/typeclasses.ml55
-rw-r--r--pretyping/typeclasses.mli10
-rw-r--r--pretyping/typeclasses_errors.ml1
-rw-r--r--pretyping/typeclasses_errors.mli1
-rw-r--r--pretyping/typing.ml204
-rw-r--r--pretyping/typing.mli9
-rw-r--r--pretyping/unification.ml380
-rw-r--r--pretyping/unification.mli32
-rw-r--r--pretyping/vnorm.ml117
-rw-r--r--pretyping/vnorm.mli4
-rw-r--r--printing/prettyp.ml17
-rw-r--r--printing/prettyp.mli8
-rw-r--r--printing/printer.ml45
-rw-r--r--printing/printer.mli13
-rw-r--r--printing/printmod.ml4
-rw-r--r--proofs/clenv.ml108
-rw-r--r--proofs/clenv.mli36
-rw-r--r--proofs/clenvtac.ml17
-rw-r--r--proofs/clenvtac.mli1
-rw-r--r--proofs/evar_refiner.ml20
-rw-r--r--proofs/goal.ml11
-rw-r--r--proofs/goal.mli12
-rw-r--r--proofs/logic.ml123
-rw-r--r--proofs/logic.mli6
-rw-r--r--proofs/pfedit.ml6
-rw-r--r--proofs/pfedit.mli12
-rw-r--r--proofs/proof.mli6
-rw-r--r--proofs/proof_global.ml5
-rw-r--r--proofs/proof_global.mli4
-rw-r--r--proofs/redexpr.ml13
-rw-r--r--proofs/redexpr.mli1
-rw-r--r--proofs/refine.ml7
-rw-r--r--proofs/refine.mli8
-rw-r--r--proofs/refiner.ml6
-rw-r--r--proofs/refiner.mli3
-rw-r--r--proofs/tacmach.ml24
-rw-r--r--proofs/tacmach.mli31
-rw-r--r--stm/stm.ml11
-rw-r--r--tactics/auto.ml70
-rw-r--r--tactics/auto.mli1
-rw-r--r--tactics/autorewrite.ml20
-rw-r--r--tactics/btermdn.ml33
-rw-r--r--tactics/btermdn.mli2
-rw-r--r--tactics/class_tactics.ml158
-rw-r--r--tactics/class_tactics.mli5
-rw-r--r--tactics/contradiction.ml27
-rw-r--r--tactics/contradiction.mli1
-rw-r--r--tactics/eauto.ml69
-rw-r--r--tactics/eauto.mli1
-rw-r--r--tactics/elim.ml18
-rw-r--r--tactics/elim.mli1
-rw-r--r--tactics/eqdecide.ml36
-rw-r--r--tactics/eqdecide.mli2
-rw-r--r--tactics/eqschemes.ml85
-rw-r--r--tactics/equality.ml218
-rw-r--r--tactics/equality.mli1
-rw-r--r--tactics/hints.ml188
-rw-r--r--tactics/hints.mli15
-rw-r--r--tactics/hipattern.ml249
-rw-r--r--tactics/hipattern.mli34
-rw-r--r--tactics/inv.ml62
-rw-r--r--tactics/inv.mli1
-rw-r--r--tactics/leminv.ml40
-rw-r--r--tactics/leminv.mli1
-rw-r--r--tactics/tacticals.ml161
-rw-r--r--tactics/tacticals.mli42
-rw-r--r--tactics/tactics.ml688
-rw-r--r--tactics/tactics.mli25
-rw-r--r--tactics/term_dnet.ml7
-rw-r--r--test-suite/output/inference.out12
-rw-r--r--test-suite/success/change_pattern.v34
-rw-r--r--test-suite/success/old_typeclass.v13
-rw-r--r--test-suite/success/rewrite_evar.v8
-rw-r--r--vernac/assumptions.ml25
-rw-r--r--vernac/auto_ind_decl.ml176
-rw-r--r--vernac/class.ml27
-rw-r--r--vernac/classes.ml29
-rw-r--r--vernac/command.ml122
-rw-r--r--vernac/command.mli2
-rw-r--r--vernac/discharge.ml2
-rw-r--r--vernac/explainErr.ml1
-rw-r--r--vernac/himsg.ml257
-rw-r--r--vernac/himsg.mli3
-rw-r--r--vernac/indschemes.ml9
-rw-r--r--vernac/lemmas.ml15
-rw-r--r--vernac/lemmas.mli6
-rw-r--r--vernac/obligations.ml19
-rw-r--r--vernac/record.ml24
-rw-r--r--vernac/search.ml34
-rw-r--r--vernac/vernacentries.ml30
240 files changed, 8497 insertions, 5898 deletions
diff --git a/CHANGES b/CHANGES
index 3b4d26598..02df8614c 100644
--- a/CHANGES
+++ b/CHANGES
@@ -6,6 +6,17 @@ Tactics
- New tactic "extensionality in H" which applies (possibly dependent)
functional extensionality in H supposed to be a quantified equality
until giving a bare equality.
+- New representation of terms that statically ensure stability by
+ evar-expansion. This has several consequences.
+ * In terms of performance, this adds a cost to every term destructuration,
+ but at the same time most eager evar normalizations were removed, which
+ couterbalances this drawback and even sometimes outperforms the old
+ implementation. For instance, many operations that would require O(n)
+ normalization of the term are now O(1) in tactics. YMMV.
+ * This triggers small changes in unification, which was not evar-insensitive.
+ Most notably, the new implementation recognizes Miller patterns that were
+ missed before because of a missing normalization step. Hopefully this should
+ be fairly uncommon.
Standard Library
diff --git a/dev/db b/dev/db
index d3503ef43..432baf8a6 100644
--- a/dev/db
+++ b/dev/db
@@ -28,6 +28,7 @@ install_printer Top_printers.pppattern
install_printer Top_printers.ppglob_constr
install_printer Top_printers.ppconstr
+install_printer Top_printers.ppeconstr
install_printer Top_printers.ppuni
install_printer Top_printers.ppuniverses
install_printer Top_printers.ppconstraints
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index cd464801b..474cc85c1 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -70,9 +70,10 @@ let ppwf_paths x = pp (Rtree.pp_tree pprecarg x)
(* term printers *)
let rawdebug = ref false
let ppevar evk = pp (str (Evd.string_of_existential evk))
-let ppconstr x = pp (Termops.print_constr x)
+let ppconstr x = pp (Termops.print_constr (EConstr.of_constr x))
+let ppeconstr x = pp (Termops.print_constr x)
let ppconstr_expr x = pp (Ppconstr.pr_constr_expr x)
-let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr x)
+let ppconstrdb x = pp(Flags.with_option rawdebug Termops.print_constr (EConstr.of_constr x))
let ppterm = ppconstr
let ppsconstr x = ppconstr (Mod_subst.force_constr x)
let ppconstr_univ x = Constrextern.with_universes ppconstr x
@@ -97,9 +98,9 @@ let ppidmap pr l = pp (pridmap pr l)
let ppevarsubst = ppidmap (fun id0 -> prset (fun (c,copt,id) ->
hov 0
- (Termops.print_constr c ++
+ (Termops.print_constr (EConstr.of_constr c) ++
(match copt with None -> mt () | Some c -> spc () ++ str "<expanded: " ++
- Termops.print_constr c ++ str">") ++
+ Termops.print_constr (EConstr.of_constr c) ++ str">") ++
(if id = id0 then mt ()
else spc () ++ str "<canonical: " ++ pr_id id ++ str ">"))))
@@ -158,15 +159,15 @@ let prdelta s = pp (Mod_subst.debug_pr_delta s)
let pp_idpred s = pp (pr_idpred s)
let pp_cpred s = pp (pr_cpred s)
let pp_transparent_state s = pp (pr_transparent_state s)
-let pp_stack_t n = pp (Reductionops.Stack.pr Termops.print_constr n)
+let pp_stack_t n = pp (Reductionops.Stack.pr (EConstr.of_constr %> Termops.print_constr) n)
let pp_cst_stack_t n = pp (Reductionops.Cst_stack.pr n)
let pp_state_t n = pp (Reductionops.pr_state n)
(* proof printers *)
let pr_evar ev = Pp.int (Evar.repr ev)
-let ppmetas metas = pp(pr_metaset metas)
-let ppevm evd = pp(pr_evar_map ~with_univs:!Flags.univ_print (Some 2) evd)
-let ppevmall evd = pp(pr_evar_map ~with_univs:!Flags.univ_print None evd)
+let ppmetas metas = pp(Termops.pr_metaset metas)
+let ppevm evd = pp(Termops.pr_evar_map ~with_univs:!Flags.univ_print (Some 2) evd)
+let ppevmall evd = pp(Termops.pr_evar_map ~with_univs:!Flags.univ_print None evd)
let pr_existentialset evars =
prlist_with_sep spc pr_evar (Evar.Set.elements evars)
let ppexistentialset evars =
@@ -177,14 +178,14 @@ let ppexistentialfilter filter = match Evd.Filter.repr filter with
let ppclenv clenv = pp(pr_clenv clenv)
let ppgoalgoal gl = pp(Goal.pr_goal gl)
let ppgoal g = pp(Printer.pr_goal g)
-let ppgoalsigma g = pp(Printer.pr_goal g ++ pr_evar_map None (Refiner.project g))
+let ppgoalsigma g = pp(Printer.pr_goal g ++ Termops.pr_evar_map None (Refiner.project g))
let pphintdb db = pp(Hints.pr_hint_db db)
let ppproofview p =
let gls,sigma = Proofview.proofview p in
- pp(pr_enum Goal.pr_goal gls ++ fnl () ++ pr_evar_map (Some 1) sigma)
+ pp(pr_enum Goal.pr_goal gls ++ fnl () ++ Termops.pr_evar_map (Some 1) sigma)
let ppopenconstr (x : Evd.open_constr) =
- let (evd,c) = x in pp (pr_evar_map (Some 2) evd ++ pr_constr c)
+ let (evd,c) = x in pp (Termops.pr_evar_map (Some 2) evd ++ pr_constr c)
(* spiwack: deactivated until a replacement is found
let pppftreestate p = pp(print_pftreestate p)
*)
@@ -215,7 +216,7 @@ let ppuniverse_context_set l = pp (pr_universe_context_set prlev l)
let ppuniverse_subst l = pp (Univ.pr_universe_subst l)
let ppuniverse_opt_subst l = pp (Universes.pr_universe_opt_subst l)
let ppuniverse_level_subst l = pp (Univ.pr_universe_level_subst l)
-let ppevar_universe_context l = pp (Evd.pr_evar_universe_context l)
+let ppevar_universe_context l = pp (Termops.pr_evar_universe_context l)
let ppconstraints c = pp (pr_constraints Level.pr c)
let ppuniverseconstraints c = pp (Universes.Constraints.pr c)
let ppuniverse_context_future c =
diff --git a/engine/eConstr.ml b/engine/eConstr.ml
new file mode 100644
index 000000000..bb9075e74
--- /dev/null
+++ b/engine/eConstr.ml
@@ -0,0 +1,782 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CSig
+open CErrors
+open Util
+open Names
+open Term
+open Context
+open Evd
+
+module API :
+sig
+module ESorts :
+sig
+type t
+val make : Sorts.t -> t
+val kind : Evd.evar_map -> t -> Sorts.t
+val unsafe_to_sorts : t -> Sorts.t
+end
+module EInstance :
+sig
+type t
+val make : Univ.Instance.t -> t
+val kind : Evd.evar_map -> t -> Univ.Instance.t
+val empty : t
+val is_empty : t -> bool
+val unsafe_to_instance : t -> Univ.Instance.t
+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 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
+val to_constr : evar_map -> t -> Constr.t
+val unsafe_to_constr : t -> Constr.t
+val unsafe_eq : (t, Constr.t) eq
+val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, t) Context.Named.Declaration.pt
+val unsafe_to_named_decl : (t, t) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt
+val unsafe_to_rel_decl : (t, t) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, t) Context.Rel.Declaration.pt
+end =
+struct
+
+module ESorts =
+struct
+ type t = Sorts.t
+ let make s = s
+ let kind sigma = function
+ | Type u -> sort_of_univ (Evd.normalize_universe sigma u)
+ | s -> s
+ let unsafe_to_sorts s = s
+end
+
+module EInstance =
+struct
+ type t = Univ.Instance.t
+ let make i = i
+ let kind sigma i =
+ if Univ.Instance.is_empty i then i
+ else Evd.normalize_universe_instance sigma i
+ let empty = Univ.Instance.empty
+ let is_empty = Univ.Instance.is_empty
+ let unsafe_to_instance t = t
+end
+
+type t = Constr.t
+
+let safe_evar_value sigma ev =
+ try Some (Evd.existential_value sigma ev)
+ with NotInstantiatedEvar | Not_found -> None
+
+let rec whd_evar sigma c =
+ match Constr.kind c with
+ | Evar ev ->
+ begin match safe_evar_value sigma ev with
+ | Some c -> whd_evar sigma c
+ | None -> c
+ end
+ | App (f, args) when isEvar f ->
+ (** Enforce smart constructor invariant on applications *)
+ let ev = 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 ->
+ (** Enforce smart constructor invariant on casts. *)
+ let ev = destEvar c0 in
+ begin match safe_evar_value sigma ev with
+ | None -> c
+ | Some c -> whd_evar sigma (mkCast (c, k, t))
+ end
+ | _ -> c
+
+let kind sigma c = Constr.kind (whd_evar sigma c)
+let kind_upto = kind
+let kind_of_type sigma c = Term.kind_of_type (whd_evar sigma c)
+let of_kind = Constr.of_kind
+let of_constr c = c
+let unsafe_to_constr c = c
+let unsafe_eq = Refl
+
+let rec to_constr sigma c = match Constr.kind c with
+| Evar ev ->
+ begin match safe_evar_value sigma ev with
+ | Some c -> to_constr sigma c
+ | None -> Constr.map (fun c -> to_constr sigma c) c
+ end
+| Sort (Type u) ->
+ let u' = Evd.normalize_universe sigma u in
+ if u' == u then c else mkSort (Sorts.sort_of_univ u')
+| Const (c', u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstU (c', u')
+| Ind (i, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkIndU (i, u')
+| Construct (co, u) when not (Univ.Instance.is_empty u) ->
+ let u' = Evd.normalize_universe_instance sigma u in
+ if u' == u then c else mkConstructU (co, u')
+| _ -> Constr.map (fun c -> to_constr sigma c) c
+
+let of_named_decl d = d
+let unsafe_to_named_decl d = d
+let of_rel_decl d = d
+let unsafe_to_rel_decl d = d
+
+end
+
+include API
+
+type types = t
+type constr = t
+type existential = t pexistential
+type fixpoint = (t, t) pfixpoint
+type cofixpoint = (t, t) pcofixpoint
+type unsafe_judgment = (constr, types) Environ.punsafe_judgment
+type unsafe_type_judgment = types Environ.punsafe_type_judgment
+type named_declaration = (constr, types) Context.Named.Declaration.pt
+type rel_declaration = (constr, types) Context.Rel.Declaration.pt
+type named_context = (constr, types) Context.Named.pt
+type rel_context = (constr, types) Context.Rel.pt
+
+let in_punivs a = (a, EInstance.empty)
+
+let mkProp = of_kind (Sort (ESorts.make Sorts.prop))
+let mkSet = of_kind (Sort (ESorts.make Sorts.set))
+let mkType u = of_kind (Sort (ESorts.make (Sorts.Type u)))
+let mkRel n = of_kind (Rel n)
+let mkVar id = of_kind (Var id)
+let mkMeta n = of_kind (Meta n)
+let mkEvar e = of_kind (Evar e)
+let mkSort s = of_kind (Sort (ESorts.make s))
+let mkCast (b, k, t) = of_kind (Cast (b, k, t))
+let mkProd (na, t, u) = of_kind (Prod (na, t, u))
+let mkLambda (na, t, c) = of_kind (Lambda (na, t, c))
+let mkLetIn (na, b, t, c) = of_kind (LetIn (na, b, t, c))
+let mkApp (f, arg) = of_kind (App (f, arg))
+let mkConstU pc = of_kind (Const pc)
+let mkConst c = of_kind (Const (in_punivs c))
+let mkIndU pi = of_kind (Ind pi)
+let mkInd i = of_kind (Ind (in_punivs i))
+let mkConstructU pc = of_kind (Construct pc)
+let mkConstruct c = of_kind (Construct (in_punivs c))
+let mkConstructUi ((ind,u),i) = of_kind (Construct ((ind,i),u))
+let mkCase (ci, c, r, p) = of_kind (Case (ci, c, r, p))
+let mkFix f = of_kind (Fix f)
+let mkCoFix f = of_kind (CoFix f)
+let mkProj (p, c) = of_kind (Proj (p, c))
+let mkArrow t1 t2 = of_kind (Prod (Anonymous, t1, t2))
+
+let applist (f, arg) = mkApp (f, Array.of_list arg)
+
+let isRel sigma c = match kind sigma c with Rel _ -> true | _ -> false
+let isVar sigma c = match kind sigma c with Var _ -> true | _ -> false
+let isInd sigma c = match kind sigma c with Ind _ -> true | _ -> false
+let isEvar sigma c = match kind sigma c with Evar _ -> true | _ -> false
+let isMeta sigma c = match kind sigma c with Meta _ -> true | _ -> false
+let isSort sigma c = match kind sigma c with Sort _ -> true | _ -> false
+let isCast sigma c = match kind sigma c with Cast _ -> true | _ -> false
+let isApp sigma c = match kind sigma c with App _ -> true | _ -> false
+let isLambda sigma c = match kind sigma c with Lambda _ -> true | _ -> false
+let isLetIn sigma c = match kind sigma c with LetIn _ -> true | _ -> false
+let isProd sigma c = match kind sigma c with Prod _ -> true | _ -> false
+let isConst sigma c = match kind sigma c with Const _ -> true | _ -> false
+let isConstruct sigma c = match kind sigma c with Construct _ -> true | _ -> false
+let isFix sigma c = match kind sigma c with Fix _ -> true | _ -> false
+let isCoFix sigma c = match kind sigma c with CoFix _ -> true | _ -> false
+let isCase sigma c = match kind sigma c with Case _ -> true | _ -> false
+let isProj sigma c = match kind sigma c with Proj _ -> true | _ -> false
+let isVarId sigma id c =
+ match kind sigma c with Var id' -> Id.equal id id' | _ -> false
+let isRelN sigma n c =
+ match kind sigma c with Rel n' -> Int.equal n n' | _ -> false
+
+let destRel sigma c = match kind sigma c with
+| Rel p -> p
+| _ -> raise DestKO
+
+let destVar sigma c = match kind sigma c with
+| Var p -> p
+| _ -> raise DestKO
+
+let destInd sigma c = match kind sigma c with
+| Ind p -> p
+| _ -> raise DestKO
+
+let destEvar sigma c = match kind sigma c with
+| Evar p -> p
+| _ -> raise DestKO
+
+let destMeta sigma c = match kind sigma c with
+| Meta p -> p
+| _ -> raise DestKO
+
+let destSort sigma c = match kind sigma c with
+| Sort p -> p
+| _ -> raise DestKO
+
+let destCast sigma c = match kind sigma c with
+| Cast (c, k, t) -> (c, k, t)
+| _ -> raise DestKO
+
+let destApp sigma c = match kind sigma c with
+| App (f, a) -> (f, a)
+| _ -> raise DestKO
+
+let destLambda sigma c = match kind sigma c with
+| Lambda (na, t, c) -> (na, t, c)
+| _ -> raise DestKO
+
+let destLetIn sigma c = match kind sigma c with
+| LetIn (na, b, t, c) -> (na, b, t, c)
+| _ -> raise DestKO
+
+let destProd sigma c = match kind sigma c with
+| Prod (na, t, c) -> (na, t, c)
+| _ -> raise DestKO
+
+let destConst sigma c = match kind sigma c with
+| Const p -> p
+| _ -> raise DestKO
+
+let destConstruct sigma c = match kind sigma c with
+| Construct p -> p
+| _ -> raise DestKO
+
+let destFix sigma c = match kind sigma c with
+| Fix p -> p
+| _ -> raise DestKO
+
+let destCoFix sigma c = match kind sigma c with
+| CoFix p -> p
+| _ -> raise DestKO
+
+let destCase sigma c = match kind sigma c with
+| Case (ci, t, c, p) -> (ci, t, c, p)
+| _ -> raise DestKO
+
+let destProj sigma c = match kind sigma c with
+| Proj (p, c) -> (p, c)
+| _ -> raise DestKO
+
+let decompose_app sigma c =
+ match kind sigma c with
+ | App (f,cl) -> (f, Array.to_list cl)
+ | _ -> (c,[])
+
+let decompose_lam sigma c =
+ let rec lamdec_rec l c = match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) c
+ | Cast (c,_,_) -> lamdec_rec l c
+ | _ -> l,c
+ in
+ lamdec_rec [] c
+
+let decompose_lam_assum sigma c =
+ let open Rel.Declaration in
+ let rec lamdec_rec l c =
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
+ | Cast (c,_,_) -> lamdec_rec l c
+ | _ -> l,c
+ in
+ lamdec_rec Context.Rel.empty c
+
+let decompose_lam_n_assum sigma n c =
+ let open Rel.Declaration in
+ if n < 0 then
+ error "decompose_lam_n_assum: integer parameter must be positive";
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) n c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_assum: not enough abstractions"
+ in
+ lamdec_rec Context.Rel.empty n c
+
+let decompose_lam_n_decls sigma n =
+ let open Rel.Declaration in
+ if n < 0 then
+ error "decompose_lam_n_decls: integer parameter must be positive";
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> lamdec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | c -> error "decompose_lam_n_decls: not enough abstractions"
+ in
+ lamdec_rec Context.Rel.empty n
+
+let lamn n env b =
+ let rec lamrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> lamrec (n-1, l, mkLambda (v,t,b))
+ | _ -> assert false
+ in
+ lamrec (n,env,b)
+
+let compose_lam l b = lamn (List.length l) l b
+
+let rec to_lambda sigma n prod =
+ if Int.equal n 0 then
+ prod
+ else
+ match kind sigma prod with
+ | Prod (na,ty,bd) -> mkLambda (na,ty,to_lambda sigma (n-1) bd)
+ | Cast (c,_,_) -> to_lambda sigma n c
+ | _ -> user_err ~hdr:"to_lambda" (Pp.mt ())
+
+let decompose_prod sigma c =
+ let rec proddec_rec l c = match kind sigma c with
+ | Prod (x,t,c) -> proddec_rec ((x,t)::l) c
+ | Cast (c,_,_) -> proddec_rec l c
+ | _ -> l,c
+ in
+ proddec_rec [] c
+
+let decompose_prod_assum sigma c =
+ let open Rel.Declaration in
+ let rec proddec_rec l c =
+ match kind sigma c with
+ | Prod (x,t,c) -> proddec_rec (Context.Rel.add (LocalAssum (x,t)) l) c
+ | LetIn (x,b,t,c) -> proddec_rec (Context.Rel.add (LocalDef (x,b,t)) l) c
+ | Cast (c,_,_) -> proddec_rec l c
+ | _ -> l,c
+ in
+ proddec_rec Context.Rel.empty c
+
+let decompose_prod_n_assum sigma n c =
+ let open Rel.Declaration in
+ if n < 0 then
+ error "decompose_prod_n_assum: integer parameter must be positive";
+ let rec prodec_rec l n c =
+ if Int.equal n 0 then l,c
+ else
+ match kind sigma c with
+ | Prod (x,t,c) -> prodec_rec (Context.Rel.add (LocalAssum (x,t)) l) (n-1) c
+ | LetIn (x,b,t,c) -> prodec_rec (Context.Rel.add (LocalDef (x,b,t)) l) (n-1) c
+ | Cast (c,_,_) -> prodec_rec l n c
+ | c -> error "decompose_prod_n_assum: not enough assumptions"
+ in
+ prodec_rec Context.Rel.empty n c
+
+let existential_type sigma (evk, args) =
+ of_constr (existential_type sigma (evk, Array.map unsafe_to_constr args))
+
+let map sigma f c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c
+ | Cast (b,k,t) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkCast (b', k, t')
+ | Prod (na,t,b) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkProd (na, t', b')
+ | Lambda (na,t,b) ->
+ let b' = f b in
+ let t' = f t in
+ if b'==b && t' == t then c
+ else mkLambda (na, t', b')
+ | LetIn (na,b,t,k) ->
+ let b' = f b in
+ let t' = f t in
+ let k' = f k in
+ if b'==b && t' == t && k'==k then c
+ else mkLetIn (na, b', t', k')
+ | App (b,l) ->
+ let b' = f b in
+ let l' = Array.smartmap f l in
+ if b'==b && l'==l then c
+ else mkApp (b', l')
+ | Proj (p,t) ->
+ let t' = f t in
+ if t' == t then c
+ else mkProj (p, t')
+ | Evar (e,l) ->
+ let l' = Array.smartmap f l in
+ if l'==l then c
+ else mkEvar (e, l')
+ | Case (ci,p,b,bl) ->
+ let b' = f b in
+ let p' = f p in
+ let bl' = Array.smartmap f bl in
+ if b'==b && p'==p && bl'==bl then c
+ else mkCase (ci, p', b', bl')
+ | Fix (ln,(lna,tl,bl)) ->
+ let tl' = Array.smartmap f tl in
+ let bl' = Array.smartmap f bl in
+ if tl'==tl && bl'==bl then c
+ else mkFix (ln,(lna,tl',bl'))
+ | CoFix(ln,(lna,tl,bl)) ->
+ let tl' = Array.smartmap f tl in
+ let bl' = Array.smartmap f bl in
+ if tl'==tl && bl'==bl then c
+ else mkCoFix (ln,(lna,tl',bl'))
+
+let map_with_binders sigma g f l c0 = match kind sigma c0 with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> c0
+ | Cast (c, k, t) ->
+ let c' = f l c in
+ let t' = f l t in
+ if c' == c && t' == t then c0
+ else mkCast (c', k, t')
+ | Prod (na, t, c) ->
+ let t' = f l t in
+ let c' = f (g l) c in
+ if t' == t && c' == c then c0
+ else mkProd (na, t', c')
+ | Lambda (na, t, c) ->
+ let t' = f l t in
+ let c' = f (g l) c in
+ if t' == t && c' == c then c0
+ else mkLambda (na, t', c')
+ | LetIn (na, b, t, c) ->
+ let b' = f l b in
+ let t' = f l t in
+ let c' = f (g l) c in
+ if b' == b && t' == t && c' == c then c0
+ else mkLetIn (na, b', t', c')
+ | App (c, al) ->
+ let c' = f l c in
+ let al' = CArray.Fun1.smartmap f l al in
+ if c' == c && al' == al then c0
+ else mkApp (c', al')
+ | Proj (p, t) ->
+ let t' = f l t in
+ if t' == t then c0
+ else mkProj (p, t')
+ | Evar (e, al) ->
+ let al' = CArray.Fun1.smartmap f l al in
+ if al' == al then c0
+ else mkEvar (e, al')
+ | Case (ci, p, c, bl) ->
+ let p' = f l p in
+ let c' = f l c in
+ let bl' = CArray.Fun1.smartmap f l bl in
+ if p' == p && c' == c && bl' == bl then c0
+ else mkCase (ci, p', c', bl')
+ | Fix (ln, (lna, tl, bl)) ->
+ let tl' = CArray.Fun1.smartmap f l tl in
+ let l' = iterate g (Array.length tl) l in
+ let bl' = CArray.Fun1.smartmap f l' bl in
+ if tl' == tl && bl' == bl then c0
+ else mkFix (ln,(lna,tl',bl'))
+ | CoFix(ln,(lna,tl,bl)) ->
+ let tl' = CArray.Fun1.smartmap f l tl in
+ let l' = iterate g (Array.length tl) l in
+ let bl' = CArray.Fun1.smartmap f l' bl in
+ mkCoFix (ln,(lna,tl',bl'))
+
+let iter sigma f c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,_,t) -> f c; f t
+ | Prod (_,t,c) -> f t; f c
+ | Lambda (_,t,c) -> f t; f c
+ | LetIn (_,b,t,c) -> f b; f t; f c
+ | App (c,l) -> f c; Array.iter f l
+ | Proj (p,c) -> f c
+ | Evar (_,l) -> Array.iter f l
+ | Case (_,p,c,bl) -> f p; f c; Array.iter f bl
+ | Fix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
+ | CoFix (_,(_,tl,bl)) -> Array.iter f tl; Array.iter f bl
+
+let iter_with_full_binders sigma g f n c =
+ let open Context.Rel.Declaration in
+ match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> ()
+ | Cast (c,_,t) -> f n c; f n t
+ | Prod (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c
+ | Lambda (na,t,c) -> f n t; f (g (LocalAssum (na, t)) n) c
+ | LetIn (na,b,t,c) -> f n b; f n t; f (g (LocalDef (na, b, t)) n) c
+ | App (c,l) -> f n c; CArray.Fun1.iter f n l
+ | Evar (_,l) -> CArray.Fun1.iter f n l
+ | Case (_,p,c,bl) -> f n p; f n c; CArray.Fun1.iter f n bl
+ | Proj (p,c) -> f n c
+ | Fix (_,(lna,tl,bl)) ->
+ Array.iter (f n) tl;
+ let n' = Array.fold_left2 (fun n na t -> g (LocalAssum (na,t)) n) n lna tl in
+ Array.iter (f n') bl
+ | CoFix (_,(lna,tl,bl)) ->
+ Array.iter (f n) tl;
+ let n' = Array.fold_left2 (fun n na t -> g (LocalAssum (na,t)) n) n lna tl in
+ Array.iter (f n') bl
+
+let iter_with_binders sigma g f n c =
+ iter_with_full_binders sigma (fun _ acc -> g acc) f n c
+
+let fold sigma f acc c = match kind sigma c with
+ | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
+ | Construct _) -> acc
+ | Cast (c,_,t) -> f (f acc c) t
+ | Prod (_,t,c) -> f (f acc t) c
+ | Lambda (_,t,c) -> f (f acc t) c
+ | LetIn (_,b,t,c) -> f (f (f acc b) t) c
+ | App (c,l) -> Array.fold_left f (f acc c) l
+ | Proj (p,c) -> f acc c
+ | Evar (_,l) -> Array.fold_left f acc l
+ | Case (_,p,c,bl) -> Array.fold_left f (f (f acc p) c) bl
+ | Fix (_,(lna,tl,bl)) ->
+ Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl
+ | CoFix (_,(lna,tl,bl)) ->
+ Array.fold_left2 (fun acc t b -> f (f acc t) b) acc tl bl
+
+let compare_gen k eq_inst eq_sort eq_constr c1 c2 =
+ (c1 == c2) || Constr.compare_head_gen_with k k eq_inst eq_sort eq_constr c1 c2
+
+let eq_constr sigma c1 c2 =
+ let kind c = kind_upto sigma c in
+ let rec eq_constr c1 c2 =
+ compare_gen kind (fun _ -> Univ.Instance.equal) Sorts.equal eq_constr c1 c2
+ in
+ eq_constr (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+let eq_constr_nounivs sigma c1 c2 =
+ let kind c = kind_upto sigma c in
+ let rec eq_constr c1 c2 =
+ compare_gen kind (fun _ _ _ -> true) (fun _ _ -> true) eq_constr c1 c2
+ in
+ eq_constr (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+let compare_constr sigma cmp c1 c2 =
+ let kind c = kind_upto sigma c in
+ let cmp c1 c2 = cmp (of_constr c1) (of_constr c2) in
+ compare_gen kind (fun _ -> Univ.Instance.equal) Sorts.equal cmp (unsafe_to_constr c1) (unsafe_to_constr c2)
+
+(** TODO: factorize with universes.ml *)
+let test_constr_universes sigma leq m n =
+ let open Universes in
+ let kind c = kind_upto sigma c in
+ if m == n then Some Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,UEq,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let leq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1,ULe,Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n = compare_gen kind eq_universes eq_sorts eq_constr' m n in
+ let res =
+ if leq then
+ let rec compare_leq m n =
+ Constr.compare_head_gen_leq_with kind kind eq_universes leq_sorts eq_constr' leq_constr' m n
+ and leq_constr' m n = m == n || compare_leq m n in
+ compare_leq m n
+ else
+ Constr.compare_head_gen_with kind kind eq_universes eq_sorts eq_constr' m n
+ in
+ if res then Some !cstrs else None
+
+let eq_constr_universes sigma m n =
+ test_constr_universes sigma false (unsafe_to_constr m) (unsafe_to_constr n)
+let leq_constr_universes sigma m n =
+ test_constr_universes sigma true (unsafe_to_constr m) (unsafe_to_constr n)
+
+let compare_head_gen_proj env sigma equ eqs eqc' m n =
+ let kind c = kind_upto sigma c in
+ match kind_upto sigma m, kind_upto sigma n with
+ | Proj (p, c), App (f, args)
+ | App (f, args), Proj (p, c) ->
+ (match kind_upto sigma 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
+ if Array.length args == npars + 1 then
+ eqc' c args.(npars)
+ else false
+ | _ -> false)
+ | _ -> Constr.compare_head_gen_with kind kind equ eqs eqc' m n
+
+let eq_constr_universes_proj env sigma m n =
+ let open Universes in
+ if m == n then Some Constraints.empty
+ else
+ let cstrs = ref Constraints.empty in
+ let eq_universes strict l l' =
+ cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in
+ let eq_sorts s1 s2 =
+ if Sorts.equal s1 s2 then true
+ else
+ (cstrs := Constraints.add
+ (Sorts.univ_of_sort s1, UEq, Sorts.univ_of_sort s2) !cstrs;
+ true)
+ in
+ let rec eq_constr' m n =
+ m == n || compare_head_gen_proj env sigma eq_universes eq_sorts eq_constr' m n
+ in
+ let res = eq_constr' (unsafe_to_constr m) (unsafe_to_constr n) in
+ if res then Some !cstrs else None
+
+module Vars =
+struct
+exception LocalOccur
+let to_constr = unsafe_to_constr
+
+(** Operations that commute with evar-normalization *)
+let lift n c = of_constr (Vars.lift n (to_constr c))
+let liftn n m c = of_constr (Vars.liftn n m (to_constr c))
+
+let substnl subst n c = of_constr (Vars.substnl (List.map to_constr subst) n (to_constr c))
+let substl subst c = of_constr (Vars.substl (List.map to_constr subst) (to_constr c))
+let subst1 c r = of_constr (Vars.subst1 (to_constr c) (to_constr r))
+
+let replace_vars subst c =
+ let map (id, c) = (id, to_constr c) in
+ of_constr (Vars.replace_vars (List.map map subst) (to_constr c))
+let substn_vars n subst c = of_constr (Vars.substn_vars n subst (to_constr c))
+let subst_vars subst c = of_constr (Vars.subst_vars subst (to_constr c))
+let subst_var subst c = of_constr (Vars.subst_var subst (to_constr c))
+
+let subst_univs_level_constr subst c =
+ of_constr (Vars.subst_univs_level_constr subst (to_constr c))
+(** Operations that dot NOT commute with evar-normalization *)
+let noccurn sigma n term =
+ let rec occur_rec n c = match kind sigma c with
+ | Rel m -> if Int.equal m n then raise LocalOccur
+ | _ -> iter_with_binders sigma succ occur_rec n c
+ in
+ try occur_rec n term; true with LocalOccur -> false
+
+let noccur_between sigma n m term =
+ let rec occur_rec n c = match kind sigma c with
+ | Rel p -> if n<=p && p<n+m then raise LocalOccur
+ | _ -> iter_with_binders sigma succ occur_rec n c
+ in
+ try occur_rec n term; true with LocalOccur -> false
+
+let closedn sigma n c =
+ let rec closed_rec n c = match kind sigma c with
+ | Rel m -> if m>n then raise LocalOccur
+ | _ -> iter_with_binders sigma succ closed_rec n c
+ in
+ try closed_rec n c; true with LocalOccur -> false
+
+let closed0 sigma c = closedn sigma 0 c
+
+let subst_of_rel_context_instance ctx subst =
+ List.map of_constr (Vars.subst_of_rel_context_instance (List.map unsafe_to_rel_decl ctx) (List.map to_constr subst))
+
+end
+
+let rec isArity sigma c =
+ match kind sigma c with
+ | Prod (_,_,c) -> isArity sigma c
+ | LetIn (_,b,_,c) -> isArity sigma (Vars.subst1 b c)
+ | Cast (c,_,_) -> isArity sigma c
+ | Sort _ -> true
+ | _ -> false
+
+let mkProd_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkProd (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkLambda_or_LetIn decl c =
+ let open Context.Rel.Declaration in
+ match decl with
+ | LocalAssum (na,t) -> mkLambda (na, t, c)
+ | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
+
+let mkNamedProd id typ c = mkProd (Name id, typ, Vars.subst_var id c)
+let mkNamedLambda id typ c = mkLambda (Name id, typ, Vars.subst_var id c)
+let mkNamedLetIn id c1 t c2 = mkLetIn (Name id, c1, t, Vars.subst_var id c2)
+
+let mkNamedProd_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedProd id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
+
+let mkNamedLambda_or_LetIn decl c =
+ let open Context.Named.Declaration in
+ match decl with
+ | LocalAssum (id,t) -> mkNamedLambda id t c
+ | LocalDef (id,b,t) -> mkNamedLetIn id b t c
+
+let it_mkProd_or_LetIn t ctx = List.fold_left (fun c d -> mkProd_or_LetIn d c) t ctx
+let it_mkLambda_or_LetIn t ctx = List.fold_left (fun c d -> mkLambda_or_LetIn d c) t ctx
+
+open Context
+open Environ
+
+let sym : type a b. (a, b) eq -> (b, a) eq = fun Refl -> Refl
+
+let cast_rel_decl :
+ type a b. (a,b) eq -> (a, a) Rel.Declaration.pt -> (b, b) Rel.Declaration.pt =
+ fun Refl x -> x
+
+let cast_rel_context :
+ type a b. (a,b) eq -> (a, a) Rel.pt -> (b, b) Rel.pt =
+ fun Refl x -> x
+
+let cast_named_decl :
+ type a b. (a,b) eq -> (a, a) Named.Declaration.pt -> (b, b) Named.Declaration.pt =
+ fun Refl x -> x
+
+let cast_named_context :
+ type a b. (a,b) eq -> (a, a) Named.pt -> (b, b) Named.pt =
+ fun Refl x -> x
+
+let push_rel d e = push_rel (cast_rel_decl unsafe_eq d) e
+let push_rel_context d e = push_rel_context (cast_rel_context unsafe_eq d) e
+let push_named d e = push_named (cast_named_decl unsafe_eq d) e
+let push_named_context d e = push_named_context (cast_named_context unsafe_eq d) e
+let push_named_context_val d e = push_named_context_val (cast_named_decl unsafe_eq d) e
+
+let rel_context e = cast_rel_context (sym unsafe_eq) (rel_context e)
+let named_context e = cast_named_context (sym unsafe_eq) (named_context e)
+
+let val_of_named_context e = val_of_named_context (cast_named_context unsafe_eq e)
+let named_context_of_val e = cast_named_context (sym unsafe_eq) (named_context_of_val e)
+
+let lookup_rel i e = cast_rel_decl (sym unsafe_eq) (lookup_rel i e)
+let lookup_named n e = cast_named_decl (sym unsafe_eq) (lookup_named n e)
+let lookup_named_val n e = cast_named_decl (sym unsafe_eq) (lookup_named_val n e)
+
+let fresh_global ?loc ?rigid ?names env sigma reference =
+ let Sigma.Sigma (t,sigma,p) =
+ Sigma.fresh_global ?loc ?rigid ?names env sigma reference in
+ Sigma.Sigma (of_constr t,sigma,p)
+
+module Unsafe =
+struct
+let to_sorts = ESorts.unsafe_to_sorts
+let to_instance = EInstance.unsafe_to_instance
+let to_constr = unsafe_to_constr
+let to_rel_decl = unsafe_to_rel_decl
+let to_named_decl = unsafe_to_named_decl
+let eq = unsafe_eq
+end
diff --git a/engine/eConstr.mli b/engine/eConstr.mli
new file mode 100644
index 000000000..3a9469e55
--- /dev/null
+++ b/engine/eConstr.mli
@@ -0,0 +1,282 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+open CSig
+open Names
+open Constr
+open Context
+open Environ
+
+type t
+(** Type of incomplete terms. Essentially a wrapper around {!Constr.t} ensuring
+ that {!Constr.kind} does not observe defined evars. *)
+
+type types = t
+type constr = t
+type existential = t pexistential
+type fixpoint = (t, t) pfixpoint
+type cofixpoint = (t, t) pcofixpoint
+type unsafe_judgment = (constr, types) Environ.punsafe_judgment
+type unsafe_type_judgment = types Environ.punsafe_type_judgment
+type named_declaration = (constr, types) Context.Named.Declaration.pt
+type rel_declaration = (constr, types) Context.Rel.Declaration.pt
+type named_context = (constr, types) Context.Named.pt
+type rel_context = (constr, types) Context.Rel.pt
+
+(** {5 Universe variables} *)
+
+module ESorts :
+sig
+ type t
+ (** Type of sorts up-to universe unification. Essentially a wrapper around
+ Sorts.t so that normalization is ensured statically. *)
+
+ val make : Sorts.t -> t
+ (** Turn a sort into an up-to sort. *)
+
+ val kind : Evd.evar_map -> t -> Sorts.t
+ (** Returns the view into the current sort. Note that the kind of a variable
+ may change if the unification state of the evar map changes. *)
+
+end
+
+module EInstance :
+sig
+ type t
+ (** Type of universe instances up-to universe unification. Similar to
+ {ESorts.t} for {Univ.Instance.t}. *)
+
+ val make : Univ.Instance.t -> t
+ val kind : Evd.evar_map -> t -> Univ.Instance.t
+ val empty : t
+ val is_empty : t -> bool
+end
+
+(** {5 Destructors} *)
+
+val kind : Evd.evar_map -> t -> (t, t, ESorts.t, EInstance.t) Constr.kind_of_term
+(** Same as {!Constr.kind} except that it expands evars and normalizes
+ universes on the fly. *)
+
+val kind_upto : Evd.evar_map -> Constr.t -> (Constr.t, Constr.t, Sorts.t, Univ.Instance.t) Constr.kind_of_term
+
+val to_constr : Evd.evar_map -> t -> Constr.t
+(** Returns the evar-normal form of the argument. See {!Evarutil.nf_evar}. *)
+
+val kind_of_type : Evd.evar_map -> t -> (t, t) Term.kind_of_type
+
+(** {5 Constructors} *)
+
+val of_kind : (t, t, ESorts.t, EInstance.t) Constr.kind_of_term -> t
+(** Construct a term from a view. *)
+
+val of_constr : Constr.t -> t
+(** Translate a kernel term into an incomplete term in O(1). *)
+
+(** {5 Insensitive primitives}
+
+ Evar-insensitive versions of the corresponding functions. See the {!Constr}
+ module for more information.
+
+*)
+
+(** {6 Constructors} *)
+
+val mkRel : int -> t
+val mkVar : Id.t -> t
+val mkMeta : metavariable -> t
+val mkEvar : t pexistential -> t
+val mkSort : Sorts.t -> t
+val mkProp : t
+val mkSet : t
+val mkType : Univ.universe -> t
+val mkCast : t * cast_kind * t -> t
+val mkProd : Name.t * t * t -> t
+val mkLambda : Name.t * t * t -> t
+val mkLetIn : Name.t * t * t * t -> t
+val mkApp : t * t array -> t
+val mkConst : constant -> t
+val mkConstU : constant * EInstance.t -> t
+val mkProj : (projection * t) -> t
+val mkInd : inductive -> t
+val mkIndU : inductive * EInstance.t -> t
+val mkConstruct : constructor -> t
+val mkConstructU : constructor * EInstance.t -> t
+val mkConstructUi : (inductive * EInstance.t) * int -> t
+val mkCase : case_info * t * t * t array -> t
+val mkFix : (t, t) pfixpoint -> t
+val mkCoFix : (t, t) pcofixpoint -> t
+val mkArrow : t -> t -> t
+
+val applist : t * t list -> t
+
+val mkProd_or_LetIn : rel_declaration -> t -> t
+val mkLambda_or_LetIn : rel_declaration -> t -> t
+val it_mkProd_or_LetIn : t -> rel_context -> t
+val it_mkLambda_or_LetIn : t -> rel_context -> t
+
+val mkNamedLambda : Id.t -> types -> constr -> constr
+val mkNamedLetIn : Id.t -> constr -> types -> constr -> constr
+val mkNamedProd : Id.t -> types -> types -> types
+val mkNamedLambda_or_LetIn : named_declaration -> types -> types
+val mkNamedProd_or_LetIn : named_declaration -> types -> types
+
+(** {6 Simple case analysis} *)
+
+val isRel : Evd.evar_map -> t -> bool
+val isVar : Evd.evar_map -> t -> bool
+val isInd : Evd.evar_map -> t -> bool
+val isEvar : Evd.evar_map -> t -> bool
+val isMeta : Evd.evar_map -> t -> bool
+val isSort : Evd.evar_map -> t -> bool
+val isCast : Evd.evar_map -> t -> bool
+val isApp : Evd.evar_map -> t -> bool
+val isLambda : Evd.evar_map -> t -> bool
+val isLetIn : Evd.evar_map -> t -> bool
+val isProd : Evd.evar_map -> t -> bool
+val isConst : Evd.evar_map -> t -> bool
+val isConstruct : Evd.evar_map -> t -> bool
+val isFix : Evd.evar_map -> t -> bool
+val isCoFix : Evd.evar_map -> t -> bool
+val isCase : Evd.evar_map -> t -> bool
+val isProj : Evd.evar_map -> t -> bool
+val isArity : Evd.evar_map -> t -> bool
+val isVarId : Evd.evar_map -> Id.t -> t -> bool
+val isRelN : Evd.evar_map -> int -> t -> bool
+
+val destRel : Evd.evar_map -> t -> int
+val destMeta : Evd.evar_map -> t -> metavariable
+val destVar : Evd.evar_map -> t -> Id.t
+val destSort : Evd.evar_map -> t -> ESorts.t
+val destCast : Evd.evar_map -> t -> t * cast_kind * t
+val destProd : Evd.evar_map -> t -> Name.t * types * types
+val destLambda : Evd.evar_map -> t -> Name.t * types * t
+val destLetIn : Evd.evar_map -> t -> Name.t * t * types * t
+val destApp : Evd.evar_map -> t -> t * t array
+val destConst : Evd.evar_map -> t -> constant * EInstance.t
+val destEvar : Evd.evar_map -> t -> t pexistential
+val destInd : Evd.evar_map -> t -> inductive * EInstance.t
+val destConstruct : Evd.evar_map -> t -> constructor * EInstance.t
+val destCase : Evd.evar_map -> t -> case_info * t * t * t array
+val destProj : Evd.evar_map -> t -> projection * t
+val destFix : Evd.evar_map -> t -> (t, t) pfixpoint
+val destCoFix : Evd.evar_map -> t -> (t, t) pcofixpoint
+
+val decompose_app : Evd.evar_map -> t -> t * t list
+
+val decompose_lam : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_lam_assum : Evd.evar_map -> t -> rel_context * t
+val decompose_lam_n_assum : Evd.evar_map -> int -> t -> rel_context * t
+val decompose_lam_n_decls : Evd.evar_map -> int -> t -> rel_context * t
+
+val compose_lam : (Name.t * t) list -> t -> t
+val to_lambda : Evd.evar_map -> int -> t -> t
+
+val decompose_prod : Evd.evar_map -> t -> (Name.t * t) list * t
+val decompose_prod_assum : Evd.evar_map -> t -> rel_context * t
+val decompose_prod_n_assum : Evd.evar_map -> int -> t -> rel_context * t
+
+val existential_type : Evd.evar_map -> existential -> types
+val whd_evar : Evd.evar_map -> constr -> constr
+
+(** {6 Equality} *)
+
+val eq_constr : Evd.evar_map -> t -> t -> bool
+val eq_constr_nounivs : Evd.evar_map -> t -> t -> bool
+val eq_constr_universes : Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val leq_constr_universes : Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val eq_constr_universes_proj : Environ.env -> Evd.evar_map -> t -> t -> Universes.universe_constraints option
+val compare_constr : Evd.evar_map -> (t -> t -> bool) -> t -> t -> bool
+
+(** {6 Iterators} *)
+
+val map : Evd.evar_map -> (t -> t) -> t -> t
+val map_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> t) -> 'a -> t -> t
+val iter : Evd.evar_map -> (t -> unit) -> t -> unit
+val iter_with_binders : Evd.evar_map -> ('a -> 'a) -> ('a -> t -> unit) -> 'a -> t -> unit
+val iter_with_full_binders : Evd.evar_map -> (rel_declaration -> 'a -> 'a) -> ('a -> t -> unit) -> 'a -> t -> unit
+val fold : Evd.evar_map -> ('a -> t -> 'a) -> 'a -> t -> 'a
+
+(** {6 Substitutions} *)
+
+module Vars :
+sig
+val lift : int -> t -> t
+val liftn : int -> int -> t -> t
+val substnl : t list -> int -> t -> t
+val substl : t list -> t -> t
+val subst1 : t -> t -> t
+
+val replace_vars : (Id.t * t) list -> t -> t
+val substn_vars : int -> Id.t list -> t -> t
+val subst_vars : Id.t list -> t -> t
+val subst_var : Id.t -> t -> t
+
+val noccurn : Evd.evar_map -> int -> t -> bool
+val noccur_between : Evd.evar_map -> int -> int -> t -> bool
+
+val closedn : Evd.evar_map -> int -> t -> bool
+val closed0 : Evd.evar_map -> t -> bool
+
+val subst_univs_level_constr : Univ.universe_level_subst -> t -> t
+val subst_of_rel_context_instance : rel_context -> t list -> t list
+
+
+end
+
+(** {5 Environment handling} *)
+
+val push_rel : rel_declaration -> env -> env
+val push_rel_context : rel_context -> env -> env
+
+val push_named : named_declaration -> env -> env
+val push_named_context : named_context -> env -> env
+val push_named_context_val : named_declaration -> named_context_val -> named_context_val
+
+val rel_context : env -> rel_context
+val named_context : env -> named_context
+
+val val_of_named_context : named_context -> named_context_val
+val named_context_of_val : named_context_val -> named_context
+
+val lookup_rel : int -> env -> rel_declaration
+val lookup_named : variable -> env -> named_declaration
+val lookup_named_val : variable -> named_context_val -> named_declaration
+
+(* XXX Missing Sigma proxy *)
+val fresh_global :
+ ?loc:Loc.t -> ?rigid:Evd.rigid -> ?names:Univ.Instance.t -> Environ.env ->
+ 'r Sigma.t -> Globnames.global_reference -> (t, 'r) Sigma.sigma
+
+(** {5 Extra} *)
+
+val of_named_decl : (Constr.t, Constr.types) Context.Named.Declaration.pt -> (t, types) Context.Named.Declaration.pt
+val of_rel_decl : (Constr.t, Constr.types) Context.Rel.Declaration.pt -> (t, types) Context.Rel.Declaration.pt
+
+(** {5 Unsafe operations} *)
+
+module Unsafe :
+sig
+ val to_constr : t -> Constr.t
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_rel_decl : (t, types) Context.Rel.Declaration.pt -> (Constr.t, Constr.types) Context.Rel.Declaration.pt
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_named_decl : (t, types) Context.Named.Declaration.pt -> (Constr.t, Constr.types) Context.Named.Declaration.pt
+ (** Physical identity. Does not care for defined evars. *)
+
+ val to_sorts : ESorts.t -> Sorts.t
+ (** Physical identity. Does not care for normalization. *)
+
+ val to_instance : EInstance.t -> Univ.Instance.t
+ (** Physical identity. Does not care for normalization. *)
+
+ val eq : (t, Constr.t) eq
+ (** Use for transparent cast between types. *)
+end
diff --git a/engine/engine.mllib b/engine/engine.mllib
index 53cbbd73e..1b670d366 100644
--- a/engine/engine.mllib
+++ b/engine/engine.mllib
@@ -1,9 +1,10 @@
Logic_monad
-Namegen
Universes
UState
Evd
Sigma
+EConstr
+Namegen
Termops
Proofview_monad
Evarutil
diff --git a/engine/evarutil.ml b/engine/evarutil.ml
index 13444cb37..1624dc93e 100644
--- a/engine/evarutil.ml
+++ b/engine/evarutil.ml
@@ -21,10 +21,6 @@ open Sigma.Notations
module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
-let safe_evar_info sigma evk =
- try Some (Evd.find sigma evk)
- with Not_found -> None
-
let safe_evar_value sigma ev =
try Some (Evd.existential_value sigma ev)
with NotInstantiatedEvar | Not_found -> None
@@ -47,10 +43,11 @@ let evd_comb2 f evdref x y =
z
let e_new_global evdref x =
- evd_comb1 (Evd.fresh_global (Global.env())) evdref x
+ EConstr.of_constr (evd_comb1 (Evd.fresh_global (Global.env())) evdref x)
let new_global evd x =
- Sigma.fresh_global (Global.env()) evd x
+ let Sigma (c, sigma, p) = Sigma.fresh_global (Global.env()) evd x in
+ Sigma (EConstr.of_constr c, sigma, p)
(****************************************************)
(* Expanding/testing/exposing existential variables *)
@@ -68,34 +65,15 @@ let rec flush_and_check_evars sigma c =
| Some c -> flush_and_check_evars sigma c)
| _ -> map_constr (flush_and_check_evars sigma) c
-(* let nf_evar_key = Profile.declare_profile "nf_evar" *)
-(* let nf_evar = Profile.profile2 nf_evar_key Reductionops.nf_evar *)
+let flush_and_check_evars sigma c =
+ flush_and_check_evars sigma (EConstr.Unsafe.to_constr c)
-let rec whd_evar sigma c =
- match kind_of_term c with
- | Evar (evk, args) ->
- begin match safe_evar_info sigma evk with
- | Some ({ evar_body = Evar_defined c } as info) ->
- let args = Array.map (fun c -> whd_evar sigma c) args in
- let c = instantiate_evar_array info c args in
- whd_evar sigma c
- | _ -> c
- end
- | Sort (Type u) ->
- let u' = Evd.normalize_universe sigma u in
- if u' == u then c else mkSort (Sorts.sort_of_univ u')
- | Const (c', u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkConstU (c', u')
- | Ind (i, u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkIndU (i, u')
- | Construct (co, u) when not (Univ.Instance.is_empty u) ->
- let u' = Evd.normalize_universe_instance sigma u in
- if u' == u then c else mkConstructU (co, u')
- | _ -> c
-
-let rec nf_evar sigma t = Constr.map (fun t -> nf_evar sigma t) (whd_evar sigma t)
+(** Term exploration up to instantiation. *)
+let kind_of_term_upto = EConstr.kind_upto
+
+let nf_evar0 sigma t = EConstr.to_constr sigma (EConstr.of_constr t)
+let whd_evar = EConstr.whd_evar
+let nf_evar sigma c = EConstr.of_constr (EConstr.to_constr sigma c)
let j_nf_evar sigma j =
{ uj_val = nf_evar sigma j.uj_val;
@@ -120,23 +98,23 @@ let e_nf_evars_and_universes evdref =
let nf_evar_map_universes evm =
let evm = Evd.nf_constraints evm in
let subst = Evd.universe_subst evm in
- if Univ.LMap.is_empty subst then evm, nf_evar evm
+ if Univ.LMap.is_empty subst then evm, nf_evar0 evm
else
let f = nf_evars_universes evm in
Evd.raw_map (fun _ -> map_evar_info f) evm, f
let nf_named_context_evar sigma ctx =
- Context.Named.map (nf_evar sigma) ctx
+ Context.Named.map (nf_evar0 sigma) ctx
let nf_rel_context_evar sigma ctx =
Context.Rel.map (nf_evar sigma) ctx
let nf_env_evar sigma env =
let nc' = nf_named_context_evar sigma (Environ.named_context env) in
- let rel' = nf_rel_context_evar sigma (Environ.rel_context env) in
- push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
+ let rel' = nf_rel_context_evar sigma (EConstr.rel_context env) in
+ EConstr.push_rel_context rel' (reset_with_named_context (val_of_named_context nc') env)
-let nf_evar_info evc info = map_evar_info (nf_evar evc) info
+let nf_evar_info evc info = map_evar_info (nf_evar0 evc) info
let nf_evar_map evm =
Evd.raw_map (fun _ evi -> nf_evar_info evm evi) evm
@@ -148,21 +126,11 @@ let nf_evar_map_undefined evm =
(* Auxiliary functions for the conversion algorithms modulo evars
*)
-(* A probably faster though more approximative variant of
- [has_undefined (nf_evar c)]: instances are not substituted and
- maybe an evar occurs in an instance and it would disappear by
- instantiation *)
-
let has_undefined_evars evd t =
let rec has_ev t =
- match kind_of_term t with
- | Evar (ev,args) ->
- (match evar_body (Evd.find evd ev) with
- | Evar_defined c ->
- has_ev c; Array.iter has_ev args
- | Evar_empty ->
- raise NotInstantiatedEvar)
- | _ -> iter_constr has_ev t in
+ match EConstr.kind evd t with
+ | Evar _ -> raise NotInstantiatedEvar
+ | _ -> EConstr.iter evd has_ev t in
try let _ = has_ev t in false
with (Not_found | NotInstantiatedEvar) -> true
@@ -171,10 +139,10 @@ let is_ground_term evd t =
let is_ground_env evd env =
let is_ground_rel_decl = function
- | RelDecl.LocalDef (_,b,_) -> is_ground_term evd b
+ | RelDecl.LocalDef (_,b,_) -> is_ground_term evd (EConstr.of_constr b)
| _ -> true in
let is_ground_named_decl = function
- | NamedDecl.LocalDef (_,b,_) -> is_ground_term evd b
+ | NamedDecl.LocalDef (_,b,_) -> is_ground_term evd (EConstr.of_constr b)
| _ -> true in
List.for_all is_ground_rel_decl (rel_context env) &&
List.for_all is_ground_named_decl (named_context env)
@@ -193,7 +161,9 @@ let is_ground_env = memo is_ground_env
exception NoHeadEvar
-let head_evar =
+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
| Evar (evk,_) -> evk
| Case (_,_,c,_) -> hrec c
@@ -202,33 +172,24 @@ let head_evar =
| Proj (p, c) -> hrec c
| _ -> raise NoHeadEvar
in
- hrec
+ hrec c
(* Expand head evar if any (currently consider only applications but I
guess it should consider Case too) *)
let whd_head_evar_stack sigma c =
- let rec whrec (c, l as s) =
- match kind_of_term c with
- | Evar (evk,args as ev) ->
- let v =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar | Not_found -> None in
- begin match v with
- | None -> s
- | Some c -> whrec (c, l)
- end
+ let rec whrec (c, l) =
+ match EConstr.kind sigma c with
| Cast (c,_,_) -> whrec (c, l)
| App (f,args) -> whrec (f, args :: l)
- | _ -> s
+ | c -> (EConstr.of_kind c, l)
in
whrec (c, [])
let whd_head_evar sigma c =
+ let open EConstr in
let (f, args) = whd_head_evar_stack sigma c in
- (** optim: don't reallocate if empty/singleton *)
match args with
- | [] -> f
| [arg] -> mkApp (f, arg)
| _ -> mkApp (f, Array.concat args)
@@ -243,7 +204,7 @@ let new_meta =
let meta_ctr = Summary.ref 0 ~name:meta_counter_summary_name in
fun () -> incr meta_ctr; !meta_ctr
-let mk_new_meta () = mkMeta(new_meta())
+let mk_new_meta () = EConstr.mkMeta(new_meta())
(* The list of non-instantiated existential declarations (order is important) *)
@@ -297,17 +258,20 @@ let make_pure_subst evi args =
*)
let csubst_subst (k, s) c =
+ (** Safe because this is a substitution *)
+ let c = EConstr.Unsafe.to_constr c in
let rec subst n c = match Constr.kind c with
| Rel m ->
if m <= n then c
- else if m - n <= k then Int.Map.find (k - m + n) s
+ else if m - n <= k then EConstr.Unsafe.to_constr (Int.Map.find (k - m + n) s)
else mkRel (m - k)
| _ -> Constr.map_with_binders succ subst n c
in
- if k = 0 then c else subst 0 c
+ let c = if k = 0 then c else subst 0 c in
+ EConstr.of_constr c
let subst2 subst vsubst c =
- csubst_subst subst (replace_vars vsubst c)
+ csubst_subst subst (EConstr.Vars.replace_vars vsubst c)
let next_ident_away id avoid =
let avoid id = Id.Set.mem id avoid in
@@ -318,24 +282,29 @@ let next_name_away na avoid =
let id = match na with Name id -> id | Anonymous -> default_non_dependent_ident in
next_ident_away_from id avoid
-type csubst = int * Constr.t Int.Map.t
+type csubst = int * EConstr.t Int.Map.t
let empty_csubst = (0, Int.Map.empty)
type ext_named_context =
- csubst * (Id.t * Constr.constr) list *
- Id.Set.t * Context.Named.t
+ csubst * (Id.t * EConstr.constr) list *
+ Id.Set.t * EConstr.named_context
let push_var id (n, s) =
- let s = Int.Map.add n (mkVar id) s in
+ let s = Int.Map.add n (EConstr.mkVar id) s in
(succ n, s)
-let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
+let push_rel_decl_to_named_context sigma decl (subst, vsubst, avoid, nc) =
+ let open EConstr in
+ let open Vars in
+ let map_decl f d =
+ NamedDecl.map_constr f d
+ in
let replace_var_named_declaration id0 id decl =
let id' = NamedDecl.get_id decl in
let id' = if Id.equal id0 id' then id else id' in
let vsubst = [id0 , mkVar id] in
- decl |> NamedDecl.set_id id' |> NamedDecl.map_constr (replace_vars vsubst)
+ decl |> NamedDecl.set_id id' |> map_decl (replace_vars vsubst)
in
let extract_if_neq id = function
| Anonymous -> None
@@ -353,7 +322,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
else
(** id_of_name_using_hdchar only depends on the rel context which is empty
here *)
- next_ident_away (id_of_name_using_hdchar empty_env (RelDecl.get_type decl) na) avoid
+ next_ident_away (id_of_name_using_hdchar empty_env sigma (RelDecl.get_type decl) na) avoid
in
match extract_if_neq id na with
| Some id0 when not (is_section_variable id0) ->
@@ -363,7 +332,7 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
context. Unless [id] is a section variable. *)
let subst = (fst subst, Int.Map.map (replace_vars [id0,mkVar id]) (snd subst)) in
let vsubst = (id0,mkVar id)::vsubst in
- let d = decl |> NamedDecl.of_rel_decl (fun _ -> id0) |> NamedDecl.map_constr (subst2 subst vsubst) in
+ let d = decl |> NamedDecl.of_rel_decl (fun _ -> id0) |> map_decl (subst2 subst vsubst) in
let nc = List.map (replace_var_named_declaration id0 id) nc in
(push_var id0 subst, vsubst, Id.Set.add id avoid, d :: nc)
| _ ->
@@ -371,12 +340,13 @@ let push_rel_decl_to_named_context decl (subst, vsubst, avoid, nc) =
incorrect. We revert to a less robust behaviour where
the new binder has name [id]. Which amounts to the same
behaviour than when [id=id0]. *)
- let d = decl |> NamedDecl.of_rel_decl (fun _ -> id) |> NamedDecl.map_constr (subst2 subst vsubst) in
+ let d = decl |> NamedDecl.of_rel_decl (fun _ -> id) |> map_decl (subst2 subst vsubst) in
(push_var id subst, vsubst, Id.Set.add id avoid, d :: nc)
-let push_rel_context_to_named_context env typ =
+let push_rel_context_to_named_context env sigma typ =
(* compute the instances relative to the named context and rel_context *)
let open Context.Named.Declaration in
+ let open EConstr in
let ids = List.map get_id (named_context env) in
let inst_vars = List.map mkVar ids in
if List.is_empty (Environ.rel_context env) then
@@ -388,7 +358,7 @@ let push_rel_context_to_named_context env typ =
(* with vars of the rel context *)
(* We do keep the instances corresponding to local definition (see above) *)
let (subst, vsubst, _, env) =
- Context.Rel.fold_outside push_rel_decl_to_named_context
+ Context.Rel.fold_outside (fun d acc -> push_rel_decl_to_named_context sigma d acc)
(rel_context env) ~init:(empty_csubst, [], avoid, named_context env) in
(val_of_named_context env, subst2 subst vsubst typ, inst_rels@inst_vars, subst, vsubst)
@@ -400,6 +370,7 @@ let default_source = (Loc.ghost,Evar_kinds.InternalHole)
let restrict_evar evd evk filter candidates =
let evd = Sigma.to_evar_map evd in
+ let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in
let evd, evk' = Evd.restrict evk filter ?candidates evd in
Sigma.Unsafe.of_pair (evk', Evd.declare_future_goal evk' evd)
@@ -410,9 +381,19 @@ let new_pure_evar_full evd evi =
Sigma.Unsafe.of_pair (evk, evd)
let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?candidates ?(store = Store.empty) ?naming ?(principal=false) typ =
+ let typ = EConstr.Unsafe.to_constr typ in
let evd = Sigma.to_evar_map evd in
+ let candidates = Option.map (fun l -> List.map EConstr.Unsafe.to_constr l) candidates in
let default_naming = Misctypes.IntroAnonymous in
let naming = Option.default default_naming naming in
+ let name = match naming with
+ | Misctypes.IntroAnonymous -> None
+ | Misctypes.IntroIdentifier id -> Some id
+ | Misctypes.IntroFresh id ->
+ let has_name id = try let _ = Evd.evar_key id evd in true with Not_found -> false in
+ let id = Namegen.next_ident_away_from id has_name in
+ Some id
+ in
let evi = {
evar_hyps = sign;
evar_concl = typ;
@@ -422,7 +403,7 @@ let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?ca
evar_candidates = candidates;
evar_extra = store; }
in
- let (evd, newevk) = Evd.new_evar evd ~naming evi in
+ let (evd, newevk) = Evd.new_evar evd ?name evi in
let evd =
if principal then Evd.declare_principal_goal newevk evd
else Evd.declare_future_goal newevk evd
@@ -430,6 +411,7 @@ let new_pure_evar sign evd ?(src=default_source) ?(filter = Filter.identity) ?ca
Sigma.Unsafe.of_pair (newevk, evd)
let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?principal instance =
+ let open EConstr in
assert (not !Flags.debug ||
List.distinct (ids_of_named_context (named_context_of_val sign)));
let Sigma (newevk, evd, p) = new_pure_evar sign evd ?src ?filter ?candidates ?store ?naming ?principal typ in
@@ -438,8 +420,9 @@ let new_evar_instance sign evd typ ?src ?filter ?candidates ?store ?naming ?prin
(* [new_evar] declares a new existential in an env env with type typ *)
(* Converting the env into the sign of the evar to define *)
let new_evar env evd ?src ?filter ?candidates ?store ?naming ?principal typ =
- let sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env typ in
- let candidates = Option.map (List.map (subst2 subst vsubst)) candidates in
+ let sign,typ',instance,subst,vsubst = push_rel_context_to_named_context env (Sigma.to_evar_map evd) typ in
+ let map c = subst2 subst vsubst c in
+ let candidates = Option.map (fun l -> List.map map l) candidates in
let instance =
match filter with
| None -> instance
@@ -453,7 +436,7 @@ let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal t
let new_type_evar env evd ?src ?filter ?naming ?principal rigid =
let Sigma (s, evd', p) = Sigma.new_sort_variable rigid evd in
- let Sigma (e, evd', q) = new_evar env evd' ?src ?filter ?naming ?principal (mkSort s) in
+ let Sigma (e, evd', q) = new_evar env evd' ?src ?filter ?naming ?principal (EConstr.mkSort s) in
Sigma ((e, s), evd', p +> q)
let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid =
@@ -464,12 +447,13 @@ let e_new_type_evar env evdref ?src ?filter ?naming ?principal rigid =
c
let new_Type ?(rigid=Evd.univ_flexible) env evd =
+ let open EConstr in
let Sigma (s, sigma, p) = Sigma.new_sort_variable rigid evd in
Sigma (mkSort s, sigma, p)
let e_new_Type ?(rigid=Evd.univ_flexible) env evdref =
let evd', s = new_sort_variable rigid !evdref in
- evdref := evd'; mkSort s
+ evdref := evd'; EConstr.mkSort s
(* The same using side-effect *)
let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?naming ?principal ty =
@@ -480,12 +464,13 @@ let e_new_evar env evdref ?(src=default_source) ?filter ?candidates ?store ?nami
(* This assumes an evar with identity instance and generalizes it over only
the De Bruijn part of the context *)
let generalize_evar_over_rels sigma (ev,args) =
+ let open EConstr in
let evi = Evd.find sigma ev in
let sign = named_context_of_val evi.evar_hyps in
List.fold_left2
(fun (c,inst as x) a d ->
- if isRel a then (mkNamedProd_or_LetIn d c,a::inst) else x)
- (evi.evar_concl,[]) (Array.to_list args) sign
+ if isRel sigma a then (mkNamedProd_or_LetIn d c,a::inst) else x)
+ (EConstr.of_constr evi.evar_concl,[]) (Array.to_list args) sign
(************************************)
(* Removing a dependency in an evar *)
@@ -523,7 +508,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
| Evar (evk,l as ev) ->
if Evd.is_defined !evdref evk then
(* If evk is already defined we replace it by its definition *)
- let nc = whd_evar !evdref c in
+ let nc = Evd.existential_value !evdref ev in
(check_and_clear_in_constr env evdref err ids global nc)
else
(* We check for dependencies to elements of ids in the
@@ -533,6 +518,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
removed *)
let evi = Evd.find_undefined !evdref evk in
let ctxt = Evd.evar_filtered_context evi in
+ let ctxt = List.map (fun d -> map_named_decl EConstr.of_constr d) ctxt in
let (rids,filter) =
List.fold_right2
(fun h a (ri,filter) ->
@@ -540,12 +526,12 @@ let rec check_and_clear_in_constr env evdref err ids global c =
(* Check if some id to clear occurs in the instance
a of rid in ev and remember the dependency *)
let check id = if Id.Set.mem id ids then raise (Depends id) in
- let () = Id.Set.iter check (collect_vars a) in
+ let () = Id.Set.iter check (collect_vars !evdref (EConstr.of_constr a)) in
(* Check if some rid to clear in the context of ev
has dependencies in another hyp of the context of ev
and transitively remember the dependency *)
let check id _ =
- if occur_var_in_decl (Global.env ()) id h
+ if occur_var_in_decl (Global.env ()) !evdref id h
then raise (Depends id)
in
let () = Id.Map.iter check ri in
@@ -578,7 +564,7 @@ let rec check_and_clear_in_constr env evdref err ids global c =
let evi' = { evi with evar_extra = extra' } in
evdref := Evd.add !evdref evk evi' ;
(* spiwack: /hacking session *)
- whd_evar !evdref c
+ Evd.existential_value !evdref ev
| _ -> map_constr (check_and_clear_in_constr env evdref err ids global) c
@@ -586,6 +572,7 @@ let clear_hyps_in_evi_main env evdref hyps terms ids =
(* clear_hyps_in_evi erases hypotheses ids in hyps, checking if some
hypothesis does not depend on a element of ids, and erases ids in
the contexts of the evars occurring in evi *)
+ let terms = List.map EConstr.Unsafe.to_constr terms in
let global = Id.Set.exists is_section_variable ids in
let terms =
List.map (check_and_clear_in_constr env evdref (OccurHypInSimpleClause None) ids global) terms in
@@ -607,7 +594,7 @@ let clear_hyps_in_evi_main env evdref hyps terms ids =
in
remove_hyps ids check_context check_value hyps
in
- (nhyps,terms)
+ (nhyps,List.map EConstr.of_constr terms)
let clear_hyps_in_evi env evdref hyps concl ids =
match clear_hyps_in_evi_main env evdref hyps [concl] ids with
@@ -696,29 +683,26 @@ let rec advance sigma evk =
let undefined_evars_of_term evd t =
let rec evrec acc c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (n, l) ->
- let acc = Array.fold_left evrec acc l in
- (try match (Evd.find evd n).evar_body with
- | Evar_empty -> Evar.Set.add n acc
- | Evar_defined c -> evrec acc c
- with Not_found -> anomaly ~label:"undefined_evars_of_term" (Pp.str "evar not found"))
- | _ -> fold_constr evrec acc c
+ let acc = Evar.Set.add n acc in
+ Array.fold_left evrec acc l
+ | _ -> EConstr.fold evd evrec acc c
in
evrec Evar.Set.empty t
let undefined_evars_of_named_context evd nc =
Context.Named.fold_outside
- (NamedDecl.fold_constr (fun c s -> Evar.Set.union s (undefined_evars_of_term evd c)))
+ (NamedDecl.fold_constr (fun c s -> Evar.Set.union s (undefined_evars_of_term evd (EConstr.of_constr c))))
nc
~init:Evar.Set.empty
let undefined_evars_of_evar_info evd evi =
- Evar.Set.union (undefined_evars_of_term evd evi.evar_concl)
+ Evar.Set.union (undefined_evars_of_term evd (EConstr.of_constr evi.evar_concl))
(Evar.Set.union
(match evi.evar_body with
| Evar_empty -> Evar.Set.empty
- | Evar_defined b -> undefined_evars_of_term evd b)
+ | Evar_defined b -> undefined_evars_of_term evd (EConstr.of_constr b))
(undefined_evars_of_named_context evd
(named_context_of_val evi.evar_hyps)))
@@ -727,6 +711,7 @@ let undefined_evars_of_evar_info evd evi =
[evar_map]. If unification only need to check superficially, tactics
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
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar e -> Option.iter occur_rec (existential_opt_value sigma e)
@@ -738,6 +723,7 @@ let occur_evar_upto sigma n c =
any type has type Type. May cause some trouble, but not so far... *)
let judge_of_new_Type evd =
+ let open EConstr in
let Sigma (s, evd', p) = Sigma.new_univ_variable univ_rigid evd in
Sigma ({ uj_val = mkSort (Type s); uj_type = mkSort (Type (Univ.super s)) }, evd', p)
@@ -748,10 +734,6 @@ let subterm_source evk (loc,k) =
(loc,Evar_kinds.SubEvar evk)
-(** Term exploration up to instantiation. *)
-let kind_of_term_upto sigma t =
- Constr.kind (whd_evar sigma t)
-
(** [eq_constr_univs_test sigma1 sigma2 t u] tests equality of [t] and
[u] up to existential variable instantiation and equalisable
universes. The term [t] is interpreted in [sigma1] while [u] is
@@ -772,5 +754,5 @@ let eq_constr_univs_test sigma1 sigma2 t u =
in
match ans with None -> false | Some _ -> true
-type type_constraint = types option
-type val_constraint = constr option
+type type_constraint = EConstr.types option
+type val_constraint = EConstr.constr option
diff --git a/engine/evarutil.mli b/engine/evarutil.mli
index 7fdc7aac7..ca9591e71 100644
--- a/engine/evarutil.mli
+++ b/engine/evarutil.mli
@@ -10,6 +10,7 @@ open Names
open Term
open Evd
open Environ
+open EConstr
(** This module provides useful higher-level functions for evar manipulation. *)
@@ -76,9 +77,9 @@ val new_evar_instance :
?principal:bool ->
constr list -> (constr, 'r) Sigma.sigma
-val make_pure_subst : evar_info -> constr array -> (Id.t * constr) list
+val make_pure_subst : evar_info -> 'a array -> (Id.t * 'a) list
-val safe_evar_value : evar_map -> existential -> constr option
+val safe_evar_value : evar_map -> Constr.existential -> Constr.constr option
(** {6 Evars/Metas switching...} *)
@@ -88,7 +89,7 @@ val non_instantiated : evar_map -> evar_info Evar.Map.t
(** [head_evar c] returns the head evar of [c] if any *)
exception NoHeadEvar
-val head_evar : constr -> existential_key (** may raise NoHeadEvar *)
+val head_evar : evar_map -> constr -> existential_key (** may raise NoHeadEvar *)
(* Expand head evar if any *)
val whd_head_evar : evar_map -> constr -> constr
@@ -128,7 +129,7 @@ val undefined_evars_of_evar_info : evar_map -> evar_info -> Evar.Set.t
(** [occur_evar_upto sigma k c] returns [true] if [k] appears in
[c]. It looks up recursively in [sigma] for the value of existential
variables. *)
-val occur_evar_upto : evar_map -> Evar.t -> Constr.t -> bool
+val occur_evar_upto : evar_map -> Evar.t -> constr -> bool
(** {6 Value/Type constraints} *)
@@ -150,7 +151,7 @@ val tj_nf_evar :
evar_map -> unsafe_type_judgment -> unsafe_type_judgment
val nf_named_context_evar : evar_map -> Context.Named.t -> Context.Named.t
-val nf_rel_context_evar : evar_map -> Context.Rel.t -> Context.Rel.t
+val nf_rel_context_evar : evar_map -> rel_context -> rel_context
val nf_env_evar : evar_map -> env -> env
val nf_evar_info : evar_map -> evar_info -> evar_info
@@ -159,39 +160,40 @@ val nf_evar_map_undefined : evar_map -> evar_map
(** Presenting terms without solved evars *)
-val nf_evars_universes : evar_map -> constr -> constr
+val nf_evars_universes : evar_map -> Constr.constr -> Constr.constr
-val nf_evars_and_universes : evar_map -> evar_map * (constr -> constr)
-val e_nf_evars_and_universes : evar_map ref -> (constr -> constr) * Universes.universe_opt_subst
+val nf_evars_and_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
+val e_nf_evars_and_universes : evar_map ref -> (Constr.constr -> Constr.constr) * Universes.universe_opt_subst
(** Normalize the evar map w.r.t. universes, after simplification of constraints.
Return the substitution function for constrs as well. *)
-val nf_evar_map_universes : evar_map -> evar_map * (constr -> constr)
+val nf_evar_map_universes : evar_map -> evar_map * (Constr.constr -> Constr.constr)
(** Replacing all evars, possibly raising [Uninstantiated_evar] *)
exception Uninstantiated_evar of existential_key
-val flush_and_check_evars : evar_map -> constr -> constr
+val flush_and_check_evars : evar_map -> constr -> Constr.constr
(** {6 Term manipulation up to instantiation} *)
(** Like {!Constr.kind} except that [kind_of_term sigma t] exposes [t]
as an evar [e] only if [e] is uninstantiated in [sigma]. Otherwise the
value of [e] in [sigma] is (recursively) used. *)
-val kind_of_term_upto : evar_map -> constr -> (constr,types) kind_of_term
+val kind_of_term_upto : evar_map -> Constr.constr ->
+ (Constr.constr, Constr.types, Sorts.t, Univ.Instance.t) kind_of_term
(** [eq_constr_univs_test sigma1 sigma2 t u] tests equality of [t] and
[u] up to existential variable instantiation and equalisable
universes. The term [t] is interpreted in [sigma1] while [u] is
interpreted in [sigma2]. The universe constraints in [sigma2] are
assumed to be an extention of those in [sigma1]. *)
-val eq_constr_univs_test : evar_map -> evar_map -> constr -> constr -> bool
+val eq_constr_univs_test : evar_map -> evar_map -> Constr.constr -> Constr.constr -> bool
(** {6 Removing hyps in evars'context}
raise OccurHypInSimpleClause if the removal breaks dependencies *)
type clear_dependency_error =
| OccurHypInSimpleClause of Id.t option
-| EvarTypingBreak of existential
+| EvarTypingBreak of Constr.existential
exception ClearDependencyError of Id.t * clear_dependency_error
@@ -208,16 +210,16 @@ val clear_hyps2_in_evi : env -> evar_map ref -> named_context_val -> types -> ty
type csubst
val empty_csubst : csubst
-val csubst_subst : csubst -> Constr.t -> Constr.t
+val csubst_subst : csubst -> constr -> constr
type ext_named_context =
- csubst * (Id.t * Constr.constr) list *
- Id.Set.t * Context.Named.t
+ csubst * (Id.t * constr) list *
+ Id.Set.t * named_context
val push_rel_decl_to_named_context :
- Context.Rel.Declaration.t -> ext_named_context -> ext_named_context
+ evar_map -> rel_declaration -> ext_named_context -> ext_named_context
-val push_rel_context_to_named_context : Environ.env -> types ->
+val push_rel_context_to_named_context : Environ.env -> evar_map -> types ->
named_context_val * types * constr list * csubst * (identifier*constr) list
val generalize_evar_over_rels : evar_map -> existential -> types * constr list
diff --git a/engine/evd.ml b/engine/evd.ml
index b7d56a698..5419a10a8 100644
--- a/engine/evd.ml
+++ b/engine/evd.ml
@@ -258,7 +258,6 @@ type evar_universe_context = UState.t
type 'a in_evar_universe_context = 'a * evar_universe_context
let empty_evar_universe_context = UState.empty
-let is_empty_evar_universe_context = UState.is_empty
let union_evar_universe_context = UState.union
let evar_universe_context_set = UState.context_set
let evar_universe_context_constraints = UState.constraints
@@ -366,7 +365,7 @@ open Misctypes
type t
val empty : t
-val add_name_undefined : intro_pattern_naming_expr -> Evar.t -> evar_info -> t -> t
+val add_name_undefined : Id.t option -> Evar.t -> evar_info -> t -> t
val remove_name_defined : Evar.t -> t -> t
val rename : Evar.t -> Id.t -> t -> t
val reassign_name_defined : Evar.t -> Evar.t -> t -> t
@@ -380,20 +379,13 @@ type t = Id.t EvMap.t * existential_key Idmap.t
let empty = (EvMap.empty, Idmap.empty)
-let add_name_newly_undefined naming evk evi (evtoid, idtoev as names) =
- let id = match naming with
- | Misctypes.IntroAnonymous -> None
- | Misctypes.IntroIdentifier id ->
- if Idmap.mem id idtoev then
- user_err (str "Already an existential evar of name " ++ pr_id id);
- Some id
- | Misctypes.IntroFresh id ->
- let id = Namegen.next_ident_away_from id (fun id -> Idmap.mem id idtoev) in
- Some id
- in
+let add_name_newly_undefined id evk evi (evtoid, idtoev as names) =
match id with
| None -> names
- | Some id -> (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
+ | Some id ->
+ if Idmap.mem id idtoev then
+ user_err (str "Already an existential evar of name " ++ pr_id id);
+ (EvMap.add evk id evtoid, Idmap.add id evk idtoev)
let add_name_undefined naming evk evi (evtoid,idtoev as evar_names) =
if EvMap.mem evk evtoid then
@@ -463,9 +455,9 @@ type evar_map = {
let rename evk id evd =
{ evd with evar_names = EvNames.rename evk id evd.evar_names }
-let add_with_name ?(naming = Misctypes.IntroAnonymous) d e i = match i.evar_body with
+let add_with_name ?name d e i = match i.evar_body with
| Evar_empty ->
- let evar_names = EvNames.add_name_undefined naming e i d.evar_names in
+ let evar_names = EvNames.add_name_undefined name e i d.evar_names in
{ d with undf_evars = EvMap.add e i d.undf_evars; evar_names }
| Evar_defined _ ->
let evar_names = EvNames.remove_name_defined e d.evar_names in
@@ -482,9 +474,9 @@ let new_untyped_evar =
let evar_ctr = Summary.ref 0 ~name:evar_counter_summary_name in
fun () -> incr evar_ctr; Evar.unsafe_of_int !evar_ctr
-let new_evar evd ?naming evi =
+let new_evar evd ?name evi =
let evk = new_untyped_evar () in
- let evd = add_with_name evd ?naming evk evi in
+ let evd = add_with_name evd ?name evk evi in
(evd, evk)
let remove d e =
@@ -505,15 +497,6 @@ let find_undefined d e = EvMap.find e d.undf_evars
let mem d e = EvMap.mem e d.undf_evars || EvMap.mem e d.defn_evars
-(* spiwack: this function loses information from the original evar_map
- it might be an idea not to export it. *)
-let to_list d =
- (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
- let l = ref [] in
- EvMap.iter (fun evk x -> l := (evk,x)::!l) d.defn_evars;
- EvMap.iter (fun evk x -> l := (evk,x)::!l) d.undf_evars;
- !l
-
let undefined_map d = d.undf_evars
let drop_all_defined d = { d with defn_evars = EvMap.empty }
@@ -732,13 +715,6 @@ let loc_of_conv_pb evd (pbty,env,t1,t2) =
(* excluding defined evars *)
-let evar_list c =
- let rec evrec acc c =
- match kind_of_term c with
- | Evar (evk, _ as ev) -> ev :: acc
- | _ -> Term.fold_constr evrec acc c in
- evrec [] c
-
let evars_of_term c =
let rec evrec acc c =
match kind_of_term c with
@@ -776,7 +752,6 @@ let evar_universe_context d = d.universes
let universe_context_set d = UState.context_set d.universes
-let pr_uctx_level = UState.pr_uctx_level
let universe_context ?names evd = UState.universe_context ?names evd.universes
let restrict_universe_context evd vars =
@@ -958,37 +933,8 @@ let universes evd = UState.ugraph evd.universes
let update_sigma_env evd env =
{ evd with universes = UState.update_sigma_env evd.universes env }
-(* Conversion w.r.t. an evar map and its local universes. *)
-
-let test_conversion_gen env evd pb t u =
- match pb with
- | Reduction.CONV ->
- Reduction.conv env
- ~evars:((existential_opt_value evd), (UState.ugraph evd.universes))
- t u
- | Reduction.CUMUL -> Reduction.conv_leq env
- ~evars:((existential_opt_value evd), (UState.ugraph evd.universes))
- t u
-
-let test_conversion env d pb t u =
- try test_conversion_gen env d pb t u; true
- with _ -> false
-
exception UniversesDiffer = UState.UniversesDiffer
-let eq_constr_univs evd t u =
- let fold cstr sigma =
- try Some (add_universe_constraints sigma cstr)
- with Univ.UniverseInconsistency _ | UniversesDiffer -> None
- in
- match Universes.eq_constr_univs_infer (UState.ugraph evd.universes) fold t u evd with
- | None -> evd, false
- | Some evd -> evd, true
-
-let e_eq_constr_univs evdref t u =
- let evd, b = eq_constr_univs !evdref t u in
- evdref := evd; b
-
(**********************************************************)
(* Side effects *)
@@ -1220,281 +1166,3 @@ module Monad =
(* Failure explanation *)
type unsolvability_explanation = SeveralInstancesFound of int
-
-(**********************************************************)
-(* Pretty-printing *)
-
-let pr_evar_suggested_name evk sigma =
- let base_id evk' evi =
- match evar_ident evk' sigma with
- | Some id -> id
- | None -> match evi.evar_source with
- | _,Evar_kinds.ImplicitArg (c,(n,Some id),b) -> id
- | _,Evar_kinds.VarInstance id -> id
- | _,Evar_kinds.GoalEvar -> Id.of_string "Goal"
- | _ ->
- let env = reset_with_named_context evi.evar_hyps (Global.env()) in
- Namegen.id_of_name_using_hdchar env evi.evar_concl Anonymous
- in
- let names = EvMap.mapi base_id sigma.undf_evars in
- let id = EvMap.find evk names in
- let fold evk' id' (seen, n) =
- if seen then (seen, n)
- else if Evar.equal evk evk' then (true, n)
- else if Id.equal id id' then (seen, succ n)
- else (seen, n)
- in
- let (_, n) = EvMap.fold fold names (false, 0) in
- if n = 0 then id else Nameops.add_suffix id (string_of_int (pred n))
-
-let pr_existential_key sigma evk = match evar_ident evk sigma with
-| None ->
- str "?" ++ pr_id (pr_evar_suggested_name evk sigma)
-| Some id ->
- str "?" ++ pr_id id
-
-let pr_instance_status (sc,typ) =
- begin match sc with
- | IsSubType -> str " [or a subtype of it]"
- | IsSuperType -> str " [or a supertype of it]"
- | Conv -> mt ()
- end ++
- begin match typ with
- | CoerceToType -> str " [up to coercion]"
- | TypeNotProcessed -> mt ()
- | TypeProcessed -> str " [type is checked]"
- end
-
-let protect f x =
- try f x
- with e -> str "EXCEPTION: " ++ str (Printexc.to_string e)
-
-let (f_print_constr, print_constr_hook) = Hook.make ()
-
-let print_constr a = protect (fun c -> Hook.get f_print_constr (Global.env ()) c) a
-
-let pr_meta_map mmap =
- let pr_name = function
- Name id -> str"[" ++ pr_id id ++ str"]"
- | _ -> mt() in
- let pr_meta_binding = function
- | (mv,Cltyp (na,b)) ->
- hov 0
- (pr_meta mv ++ pr_name na ++ str " : " ++
- print_constr b.rebus ++ fnl ())
- | (mv,Clval(na,(b,s),t)) ->
- hov 0
- (pr_meta mv ++ pr_name na ++ str " := " ++
- print_constr b.rebus ++
- str " : " ++ print_constr t.rebus ++
- spc () ++ pr_instance_status s ++ fnl ())
- in
- prlist pr_meta_binding (metamap_to_list mmap)
-
-let pr_decl (decl,ok) =
- match decl with
- | LocalAssum (id,_) -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
- | LocalDef (id,c,_) -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
- print_constr c ++ str (if ok then ")" else "}")
-
-let pr_evar_source = function
- | Evar_kinds.NamedHole id -> pr_id id
- | Evar_kinds.QuestionMark _ -> str "underscore"
- | Evar_kinds.CasesType false -> str "pattern-matching return predicate"
- | Evar_kinds.CasesType true ->
- str "subterm of pattern-matching return predicate"
- | Evar_kinds.BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
- | Evar_kinds.BinderType Anonymous -> str "type of anonymous binder"
- | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
- let id = Option.get ido in
- str "parameter " ++ pr_id id ++ spc () ++ str "of" ++
- spc () ++ print_constr (printable_constr_of_global c)
- | Evar_kinds.InternalHole -> str "internal placeholder"
- | Evar_kinds.TomatchTypeParameter (ind,n) ->
- pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind)
- | Evar_kinds.GoalEvar -> str "goal evar"
- | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause"
- | Evar_kinds.MatchingVar _ -> str "matching variable"
- | Evar_kinds.VarInstance id -> str "instance of " ++ pr_id id
- | Evar_kinds.SubEvar evk ->
- str "subterm of " ++ str (string_of_existential evk)
-
-let pr_evar_info evi =
- let phyps =
- try
- let decls = match Filter.repr (evar_filter evi) with
- | None -> List.map (fun c -> (c, true)) (evar_context evi)
- | Some filter -> List.combine (evar_context evi) filter
- in
- prlist_with_sep spc pr_decl (List.rev decls)
- with Invalid_argument _ -> str "Ill-formed filtered context" in
- let pty = print_constr evi.evar_concl in
- let pb =
- match evi.evar_body with
- | Evar_empty -> mt ()
- | Evar_defined c -> spc() ++ str"=> " ++ print_constr c
- in
- let candidates =
- match evi.evar_body, evi.evar_candidates with
- | Evar_empty, Some l ->
- spc () ++ str "{" ++
- prlist_with_sep (fun () -> str "|") print_constr l ++ str "}"
- | _ ->
- mt ()
- in
- let src = str "(" ++ pr_evar_source (snd evi.evar_source) ++ str ")" in
- hov 2
- (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
- candidates ++ spc() ++ src)
-
-let compute_evar_dependency_graph (sigma : evar_map) =
- (* Compute the map binding ev to the evars whose body depends on ev *)
- let fold evk evi acc =
- let fold_ev evk' acc =
- let tab =
- try EvMap.find evk' acc
- with Not_found -> Evar.Set.empty
- in
- EvMap.add evk' (Evar.Set.add evk tab) acc
- in
- match evar_body evi with
- | Evar_empty -> assert false
- | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc
- in
- EvMap.fold fold sigma.defn_evars EvMap.empty
-
-let evar_dependency_closure n sigma =
- (** Create the DAG of depth [n] representing the recursive dependencies of
- undefined evars. *)
- let graph = compute_evar_dependency_graph sigma in
- let rec aux n curr accu =
- if Int.equal n 0 then Evar.Set.union curr accu
- else
- let fold evk accu =
- try
- let deps = EvMap.find evk graph in
- Evar.Set.union deps accu
- with Not_found -> accu
- in
- (** Consider only the newly added evars *)
- let ncurr = Evar.Set.fold fold curr Evar.Set.empty in
- (** Merge the others *)
- let accu = Evar.Set.union curr accu in
- aux (n - 1) ncurr accu
- in
- let undef = EvMap.domain (undefined_map sigma) in
- aux n undef Evar.Set.empty
-
-let evar_dependency_closure n sigma =
- let deps = evar_dependency_closure n sigma in
- let map = EvMap.bind (fun ev -> find sigma ev) deps in
- EvMap.bindings map
-
-let has_no_evar sigma =
- EvMap.is_empty sigma.defn_evars && EvMap.is_empty sigma.undf_evars
-
-let pr_evd_level evd = pr_uctx_level evd.universes
-
-let pr_evar_universe_context ctx =
- let prl = pr_uctx_level ctx in
- if is_empty_evar_universe_context ctx then mt ()
- else
- (str"UNIVERSES:"++brk(0,1)++
- h 0 (Univ.pr_universe_context_set prl (evar_universe_context_set ctx)) ++ fnl () ++
- str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
- h 0 (Univ.LSet.pr prl (UState.algebraics ctx)) ++ fnl() ++
- str"UNDEFINED UNIVERSES:"++brk(0,1)++
- h 0 (Universes.pr_universe_opt_subst (UState.subst ctx)) ++ fnl())
-
-let print_env_short env =
- let pr_rel_decl = function
- | RelDecl.LocalAssum (n,_) -> pr_name n
- | RelDecl.LocalDef (n,b,_) -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")"
- in
- let pr_named_decl = NamedDecl.to_rel_decl %> pr_rel_decl in
- let nc = List.rev (named_context env) in
- let rc = List.rev (rel_context env) in
- str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
- str "[" ++ pr_sequence pr_rel_decl rc ++ str "]"
-
-let pr_evar_constraints pbs =
- let pr_evconstr (pbty, env, t1, t2) =
- let env =
- (** We currently allow evar instances to refer to anonymous de
- Bruijn indices, so we protect the error printing code in this
- case by giving names to every de Bruijn variable in the
- rel_context of the conversion problem. MS: we should rather
- stop depending on anonymous variables, they can be used to
- indicate independency. Also, this depends on a strategy for
- naming/renaming. *)
- Namegen.make_all_name_different env
- in
- print_env_short env ++ spc () ++ str "|-" ++ spc () ++
- Hook.get f_print_constr env t1 ++ spc () ++
- str (match pbty with
- | Reduction.CONV -> "=="
- | Reduction.CUMUL -> "<=") ++
- spc () ++ Hook.get f_print_constr env t2
- in
- prlist_with_sep fnl pr_evconstr pbs
-
-let pr_evar_map_gen with_univs pr_evars sigma =
- let { universes = uvs } = sigma in
- let evs = if has_no_evar sigma then mt () else pr_evars sigma ++ fnl ()
- and svs = if with_univs then pr_evar_universe_context uvs else mt ()
- and cstrs =
- if List.is_empty sigma.conv_pbs then mt ()
- else
- str "CONSTRAINTS:" ++ brk (0, 1) ++
- pr_evar_constraints sigma.conv_pbs ++ fnl ()
- and metas =
- if Metamap.is_empty sigma.metas then mt ()
- else
- str "METAS:" ++ brk (0, 1) ++ pr_meta_map sigma.metas
- in
- evs ++ svs ++ cstrs ++ metas
-
-let pr_evar_list sigma l =
- let pr (ev, evi) =
- h 0 (str (string_of_existential ev) ++
- str "==" ++ pr_evar_info evi ++
- (if evi.evar_body == Evar_empty
- then str " {" ++ pr_existential_key sigma ev ++ str "}"
- else mt ()))
- in
- h 0 (prlist_with_sep fnl pr l)
-
-let pr_evar_by_depth depth sigma = match depth with
-| None ->
- (* Print all evars *)
- str"EVARS:"++brk(0,1)++pr_evar_list sigma (to_list sigma)++fnl()
-| Some n ->
- (* Print all evars *)
- str"UNDEFINED EVARS:"++
- (if Int.equal n 0 then mt() else str" (+level "++int n++str" closure):")++
- brk(0,1)++
- pr_evar_list sigma (evar_dependency_closure n sigma)++fnl()
-
-let pr_evar_by_filter filter sigma =
- let defined = Evar.Map.filter filter sigma.defn_evars in
- let undefined = Evar.Map.filter filter sigma.undf_evars in
- let prdef =
- if Evar.Map.is_empty defined then mt ()
- else str "DEFINED EVARS:" ++ brk (0, 1) ++
- pr_evar_list sigma (Evar.Map.bindings defined)
- in
- let prundef =
- if Evar.Map.is_empty undefined then mt ()
- else str "UNDEFINED EVARS:" ++ brk (0, 1) ++
- pr_evar_list sigma (Evar.Map.bindings undefined)
- in
- prdef ++ prundef
-
-let pr_evar_map ?(with_univs=true) depth sigma =
- pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_depth depth sigma) sigma
-
-let pr_evar_map_filter ?(with_univs=true) filter sigma =
- pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_filter filter sigma) sigma
-
-let pr_metaset metas =
- str "[" ++ pr_sequence pr_meta (Metaset.elements metas) ++ str "]"
diff --git a/engine/evd.mli b/engine/evd.mli
index 5619b7af2..9c40c8b71 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -149,7 +149,7 @@ val has_undefined : evar_map -> bool
there are uninstantiated evars in [sigma]. *)
val new_evar : evar_map ->
- ?naming:Misctypes.intro_pattern_naming_expr -> evar_info -> evar_map * evar
+ ?name:Id.t -> evar_info -> evar_map * evar
(** Creates a fresh evar mapping to the given information. *)
val add : evar_map -> evar -> evar_info -> evar_map
@@ -420,10 +420,6 @@ val loc_of_conv_pb : evar_map -> evar_constraint -> Loc.t
contained in the object; need the term to be evar-normal otherwise
defined evars are returned too. *)
-val evar_list : constr -> existential list
- (** excluding evars in instances of evars and collected with
- redundancies from right to left (used by tactic "instantiate") *)
-
val evars_of_term : constr -> Evar.Set.t
(** including evars in instances of evars *)
@@ -584,19 +580,6 @@ val fresh_constructor_instance : ?loc:Loc.t -> env -> evar_map -> constructor ->
val fresh_global : ?loc:Loc.t -> ?rigid:rigid -> ?names:Univ.Instance.t -> env ->
evar_map -> Globnames.global_reference -> evar_map * constr
-(********************************************************************
- Conversion w.r.t. an evar map, not unifying universes. See
- [Reductionops.infer_conv] for conversion up-to universes. *)
-
-val test_conversion : env -> evar_map -> conv_pb -> constr -> constr -> bool
-(** WARNING: This does not allow unification of universes *)
-
-val eq_constr_univs : evar_map -> constr -> constr -> evar_map * bool
-(** Syntactic equality up to universes, recording the associated constraints *)
-
-val e_eq_constr_univs : evar_map ref -> constr -> constr -> bool
-(** Syntactic equality up to universes. *)
-
(********************************************************************)
(* constr with holes and pending resolution of classes, conversion *)
(* problems, candidates, etc. *)
@@ -608,22 +591,6 @@ type open_constr = evar_map * constr (* Special case when before is empty *)
type unsolvability_explanation = SeveralInstancesFound of int
(** Failure explanation. *)
-val pr_existential_key : evar_map -> evar -> Pp.std_ppcmds
-
-val pr_evar_suggested_name : existential_key -> evar_map -> Id.t
-
-(** {5 Debug pretty-printers} *)
-
-val print_constr_hook : (Environ.env -> constr -> Pp.std_ppcmds) Hook.t
-val pr_evar_info : evar_info -> Pp.std_ppcmds
-val pr_evar_constraints : evar_constraint list -> Pp.std_ppcmds
-val pr_evar_map : ?with_univs:bool -> int option -> evar_map -> Pp.std_ppcmds
-val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> evar_info -> bool) ->
- evar_map -> Pp.std_ppcmds
-val pr_metaset : Metaset.t -> Pp.std_ppcmds
-val pr_evar_universe_context : evar_universe_context -> Pp.std_ppcmds
-val pr_evd_level : evar_map -> Univ.Level.t -> Pp.std_ppcmds
-
(** {5 Deprecated functions} *)
val create_evar_defs : evar_map -> evar_map
diff --git a/engine/namegen.ml b/engine/namegen.ml
index e56ec2877..3b979f206 100644
--- a/engine/namegen.ml
+++ b/engine/namegen.ml
@@ -15,12 +15,13 @@
open Util
open Names
open Term
+open Environ
+open EConstr
open Vars
open Nametab
open Nameops
open Libnames
open Globnames
-open Environ
open Context.Rel.Declaration
module RelDecl = Context.Rel.Declaration
@@ -84,13 +85,20 @@ let is_section_variable id =
(**********************************************************************)
(* Generating "intuitive" names from its type *)
-let head_name c = (* Find the head constant of a constr if any *)
+let global_of_constr = function
+| Const (c, _) -> ConstRef c
+| Ind (i, _) -> IndRef i
+| Construct (c, _) -> ConstructRef c
+| Var id -> VarRef id
+| _ -> assert false
+
+let head_name sigma c = (* Find the head constant of a constr if any *)
let rec hdrec c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c)
| Cast (c,_,_) | App (c,_) -> hdrec c
| Proj (kn,_) -> Some (Label.to_id (con_label (Projection.constant kn)))
- | Const _ | Ind _ | Construct _ | Var _ ->
+ | Const _ | Ind _ | Construct _ | Var _ as c ->
Some (basename_of_global (global_of_constr c))
| Fix ((_,i),(lna,_,_)) | CoFix (i,(lna,_,_)) ->
Some (match lna.(i) with Name id -> id | _ -> assert false)
@@ -105,9 +113,9 @@ let sort_hdchar = function
| Prop(_) -> "P"
| Type(_) -> "T"
-let hdchar env c =
+let hdchar env sigma c =
let rec hdrec k c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (_,_,c) | Lambda (_,_,c) | LetIn (_,_,_,c) -> hdrec (k+1) c
| Cast (c,_,_) | App (c,_) -> hdrec k c
| Proj (kn,_) -> lowercase_first_char (Label.to_id (con_label (Projection.constant kn)))
@@ -115,11 +123,11 @@ let hdchar env c =
| Ind (x,_) -> lowercase_first_char (basename_of_global (IndRef x))
| Construct (x,_) -> lowercase_first_char (basename_of_global (ConstructRef x))
| Var id -> lowercase_first_char id
- | Sort s -> sort_hdchar s
+ | Sort s -> sort_hdchar (ESorts.kind sigma s)
| Rel n ->
(if n<=k then "p" (* the initial term is flexible product/function *)
else
- try match Environ.lookup_rel (n-k) env with
+ try match lookup_rel (n-k) env with
| LocalAssum (Name id,_) | LocalDef (Name id,_,_) -> lowercase_first_char id
| LocalAssum (Anonymous,t) | LocalDef (Anonymous,_,t) -> hdrec 0 (lift (n-k) t)
with Not_found -> "y")
@@ -131,41 +139,41 @@ let hdchar env c =
in
hdrec 0 c
-let id_of_name_using_hdchar env a = function
- | Anonymous -> Id.of_string (hdchar env a)
+let id_of_name_using_hdchar env sigma a = function
+ | Anonymous -> Id.of_string (hdchar env sigma a)
| Name id -> id
-let named_hd env a = function
- | Anonymous -> Name (Id.of_string (hdchar env a))
+let named_hd env sigma a = function
+ | Anonymous -> Name (Id.of_string (hdchar env sigma a))
| x -> x
-let mkProd_name env (n,a,b) = mkProd (named_hd env a n, a, b)
-let mkLambda_name env (n,a,b) = mkLambda (named_hd env a n, a, b)
+let mkProd_name env sigma (n,a,b) = mkProd (named_hd env sigma a n, a, b)
+let mkLambda_name env sigma (n,a,b) = mkLambda (named_hd env sigma a n, a, b)
let lambda_name = mkLambda_name
let prod_name = mkProd_name
-let prod_create env (a,b) = mkProd (named_hd env a Anonymous, a, b)
-let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous, a, b)
+let prod_create env sigma (a,b) = mkProd (named_hd env sigma a Anonymous, a, b)
+let lambda_create env sigma (a,b) = mkLambda (named_hd env sigma a Anonymous, a, b)
-let name_assumption env = function
- | LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
- | LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+let name_assumption env sigma = function
+ | LocalAssum (na,t) -> LocalAssum (named_hd env sigma t na, t)
+ | LocalDef (na,c,t) -> LocalDef (named_hd env sigma c na, c, t)
-let name_context env hyps =
+let name_context env sigma hyps =
snd
(List.fold_left
(fun (env,hyps) d ->
- let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
+ let d' = name_assumption env sigma d in (push_rel d' env, d' :: hyps))
(env,[]) (List.rev hyps))
-let mkProd_or_LetIn_name env b d = mkProd_or_LetIn (name_assumption env d) b
-let mkLambda_or_LetIn_name env b d = mkLambda_or_LetIn (name_assumption env d)b
+let mkProd_or_LetIn_name env sigma b d = mkProd_or_LetIn (name_assumption env sigma d) b
+let mkLambda_or_LetIn_name env sigma b d = mkLambda_or_LetIn (name_assumption env sigma d) b
-let it_mkProd_or_LetIn_name env b hyps =
- it_mkProd_or_LetIn b (name_context env hyps)
-let it_mkLambda_or_LetIn_name env b hyps =
- it_mkLambda_or_LetIn b (name_context env hyps)
+let it_mkProd_or_LetIn_name env sigma b hyps =
+ it_mkProd_or_LetIn b (name_context env sigma hyps)
+let it_mkLambda_or_LetIn_name env sigma b hyps =
+ it_mkLambda_or_LetIn b (name_context env sigma hyps)
(**********************************************************************)
(* Fresh names *)
@@ -185,10 +193,10 @@ let restart_subscript id =
*** make_ident id (Some 0) *** but compatibility would be lost... *)
forget_subscript id
-let visible_ids (nenv, c) =
+let visible_ids sigma (nenv, c) =
let accu = ref (Refset_env.empty, Int.Set.empty, Id.Set.empty) in
- let rec visible_ids n c = match kind_of_term c with
- | Const _ | Ind _ | Construct _ | Var _ ->
+ let rec visible_ids n c = match EConstr.kind sigma c with
+ | Const _ | Ind _ | Construct _ | Var _ as c ->
let (gseen, vseen, ids) = !accu in
let g = global_of_constr c in
if not (Refset_env.mem g gseen) then
@@ -218,7 +226,7 @@ let visible_ids (nenv, c) =
| _ -> ids
in
accu := (gseen, vseen, ids)
- | _ -> Constr.iter_with_binders succ visible_ids n c
+ | _ -> EConstr.iter_with_binders sigma succ visible_ids n c
in
let () = visible_ids 1 c in
let (_, _, ids) = !accu in
@@ -228,9 +236,9 @@ let visible_ids (nenv, c) =
(* 1- Looks for a fresh name for printing in cases pattern *)
-let next_name_away_in_cases_pattern env_t na avoid =
+let next_name_away_in_cases_pattern sigma env_t na avoid =
let id = match na with Name id -> id | Anonymous -> default_dependent_ident in
- let visible = visible_ids env_t in
+ let visible = visible_ids sigma env_t in
let bad id = Id.List.mem id avoid || is_constructor id
|| Id.Set.mem id visible in
next_ident_away_from id bad
@@ -292,7 +300,7 @@ let next_name_away_with_default_using_types default na avoid t =
let next_name_away = next_name_away_with_default default_non_dependent_string
-let make_all_name_different env =
+let make_all_name_different env sigma =
(** FIXME: this is inefficient, but only used in printing *)
let avoid = ref (Id.Set.elements (Context.Named.to_vars (named_context env))) in
let sign = named_context_val env in
@@ -300,7 +308,7 @@ let make_all_name_different env =
let env0 = reset_with_named_context sign env in
Context.Rel.fold_outside
(fun decl newenv ->
- let na = named_hd newenv (RelDecl.get_type decl) (RelDecl.get_name decl) in
+ let na = named_hd newenv sigma (RelDecl.get_type decl) (RelDecl.get_name decl) in
let id = next_name_away na !avoid in
avoid := id::!avoid;
push_rel (RelDecl.set_name (Name id) decl) newenv)
@@ -311,12 +319,12 @@ let make_all_name_different env =
looks for name of same base with lower available subscript beyond current
subscript *)
-let next_ident_away_for_default_printing env_t id avoid =
- let visible = visible_ids env_t in
+let next_ident_away_for_default_printing sigma env_t id avoid =
+ let visible = visible_ids sigma env_t in
let bad id = Id.List.mem id avoid || Id.Set.mem id visible in
next_ident_away_from id bad
-let next_name_away_for_default_printing env_t na avoid =
+let next_name_away_for_default_printing sigma env_t na avoid =
let id = match na with
| Name id -> id
| Anonymous ->
@@ -324,7 +332,7 @@ let next_name_away_for_default_printing env_t na avoid =
(* taken into account by the function compute_displayed_name_in; *)
(* just in case, invent a valid name *)
default_non_dependent_ident in
- next_ident_away_for_default_printing env_t id avoid
+ next_ident_away_for_default_printing sigma env_t id avoid
(**********************************************************************)
(* Displaying terms avoiding bound variables clashes *)
@@ -349,45 +357,45 @@ type renaming_flags =
| RenamingForGoal
| RenamingElsewhereFor of (Name.t list * constr)
-let next_name_for_display flags =
+let next_name_for_display sigma flags =
match flags with
- | RenamingForCasesPattern env_t -> next_name_away_in_cases_pattern env_t
+ | RenamingForCasesPattern env_t -> next_name_away_in_cases_pattern sigma env_t
| RenamingForGoal -> next_name_away_in_goal
- | RenamingElsewhereFor env_t -> next_name_away_for_default_printing env_t
+ | RenamingElsewhereFor env_t -> next_name_away_for_default_printing sigma env_t
(* Remark: Anonymous var may be dependent in Evar's contexts *)
-let compute_displayed_name_in flags avoid na c =
+let compute_displayed_name_in sigma flags avoid na c =
match na with
- | Anonymous when noccurn 1 c ->
+ | Anonymous when noccurn sigma 1 c ->
(Anonymous,avoid)
| _ ->
- let fresh_id = next_name_for_display flags na avoid in
- let idopt = if noccurn 1 c then Anonymous else Name fresh_id in
+ let fresh_id = next_name_for_display sigma flags na avoid in
+ let idopt = if noccurn sigma 1 c then Anonymous else Name fresh_id in
(idopt, fresh_id::avoid)
-let compute_and_force_displayed_name_in flags avoid na c =
+let compute_and_force_displayed_name_in sigma flags avoid na c =
match na with
- | Anonymous when noccurn 1 c ->
+ | Anonymous when noccurn sigma 1 c ->
(Anonymous,avoid)
| _ ->
- let fresh_id = next_name_for_display flags na avoid in
+ let fresh_id = next_name_for_display sigma flags na avoid in
(Name fresh_id, fresh_id::avoid)
-let compute_displayed_let_name_in flags avoid na c =
- let fresh_id = next_name_for_display flags na avoid in
+let compute_displayed_let_name_in sigma flags avoid na c =
+ let fresh_id = next_name_for_display sigma flags na avoid in
(Name fresh_id, fresh_id::avoid)
-let rename_bound_vars_as_displayed avoid env c =
+let rename_bound_vars_as_displayed sigma avoid env c =
let rec rename avoid env c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (na,c1,c2) ->
let na',avoid' =
- compute_displayed_name_in
+ compute_displayed_name_in sigma
(RenamingElsewhereFor (env,c2)) avoid na c2 in
mkProd (na', c1, rename avoid' (na' :: env) c2)
| LetIn (na,c1,t,c2) ->
let na',avoid' =
- compute_displayed_let_name_in
+ compute_displayed_let_name_in sigma
(RenamingElsewhereFor (env,c2)) avoid na c2 in
mkLetIn (na',c1,t, rename avoid' (na' :: env) c2)
| Cast (c,k,t) -> mkCast (rename avoid env c, k,t)
diff --git a/engine/namegen.mli b/engine/namegen.mli
index 33ce9a34d..058b1c228 100644
--- a/engine/namegen.mli
+++ b/engine/namegen.mli
@@ -11,6 +11,8 @@
open Names
open Term
open Environ
+open Evd
+open EConstr
(*********************************************************************
Conventional default names *)
@@ -26,27 +28,27 @@ val default_dependent_ident : Id.t (* "x" *)
val lowercase_first_char : Id.t -> string
val sort_hdchar : sorts -> string
-val hdchar : env -> types -> string
-val id_of_name_using_hdchar : env -> types -> Name.t -> Id.t
-val named_hd : env -> types -> Name.t -> Name.t
-val head_name : types -> Id.t option
+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
+val head_name : evar_map -> types -> Id.t option
-val mkProd_name : env -> Name.t * types * types -> types
-val mkLambda_name : env -> Name.t * types * constr -> constr
+val mkProd_name : env -> evar_map -> Name.t * types * types -> types
+val mkLambda_name : env -> evar_map -> Name.t * types * constr -> constr
(** Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
-val prod_name : env -> Name.t * types * types -> types
-val lambda_name : env -> Name.t * types * constr -> constr
+val prod_name : env -> evar_map -> Name.t * types * types -> types
+val lambda_name : env -> evar_map -> Name.t * types * constr -> constr
-val prod_create : env -> types * types -> constr
-val lambda_create : env -> types * constr -> constr
-val name_assumption : env -> Context.Rel.Declaration.t -> Context.Rel.Declaration.t
-val name_context : env -> Context.Rel.t -> Context.Rel.t
+val prod_create : env -> evar_map -> types * types -> constr
+val lambda_create : env -> evar_map -> types * constr -> constr
+val name_assumption : env -> evar_map -> rel_declaration -> rel_declaration
+val name_context : env -> evar_map -> rel_context -> rel_context
-val mkProd_or_LetIn_name : env -> types -> Context.Rel.Declaration.t -> types
-val mkLambda_or_LetIn_name : env -> constr -> Context.Rel.Declaration.t -> constr
-val it_mkProd_or_LetIn_name : env -> types -> Context.Rel.t -> types
-val it_mkLambda_or_LetIn_name : env -> constr -> Context.Rel.t -> constr
+val mkProd_or_LetIn_name : env -> evar_map -> types -> rel_declaration -> types
+val mkLambda_or_LetIn_name : env -> evar_map -> constr -> rel_declaration -> constr
+val it_mkProd_or_LetIn_name : env -> evar_map -> types -> rel_context -> types
+val it_mkLambda_or_LetIn_name : env -> evar_map -> constr -> rel_context -> constr
(*********************************************************************
Fresh names *)
@@ -98,16 +100,16 @@ type renaming_flags =
| RenamingForGoal (** avoid all globals (as in intro) *)
| RenamingElsewhereFor of (Name.t list * constr)
-val make_all_name_different : env -> env
+val make_all_name_different : env -> evar_map -> env
val compute_displayed_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_and_force_displayed_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val compute_displayed_let_name_in :
- renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
+ evar_map -> renaming_flags -> Id.t list -> Name.t -> constr -> Name.t * Id.t list
val rename_bound_vars_as_displayed :
- Id.t list -> Name.t list -> types -> types
+ evar_map -> Id.t list -> Name.t list -> types -> types
(**********************************************************************)
(* Naming strategy for arguments in Prop when eliminating inductive types *)
diff --git a/engine/proofview.ml b/engine/proofview.ml
index 721389af4..9c264439b 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -24,7 +24,7 @@ type proofview = Proofview_monad.proofview
(* The first items in pairs below are proofs (under construction).
The second items in the pairs below are statements that are being proved. *)
-type entry = (Term.constr * Term.types) list
+type entry = (EConstr.constr * EConstr.types) list
(** Returns a stylised view of a proofview for use by, for instance,
ide-s. *)
@@ -37,15 +37,16 @@ let proofview p =
p.comb , p.solution
let compact el ({ solution } as pv) =
- let nf = Evarutil.nf_evar solution in
+ let nf c = Evarutil.nf_evar solution c in
+ let nf0 c = EConstr.Unsafe.to_constr (Evarutil.nf_evar solution (EConstr.of_constr c)) in
let size = Evd.fold (fun _ _ i -> i+1) solution 0 in
let new_el = List.map (fun (t,ty) -> nf t, nf ty) el in
let pruned_solution = Evd.drop_all_defined solution in
let apply_subst_einfo _ ei =
Evd.({ ei with
- evar_concl = nf ei.evar_concl;
- evar_hyps = Environ.map_named_val nf ei.evar_hyps;
- evar_candidates = Option.map (List.map nf) ei.evar_candidates }) in
+ evar_concl = nf0 ei.evar_concl;
+ evar_hyps = Environ.map_named_val nf0 ei.evar_hyps;
+ evar_candidates = Option.map (List.map nf0) ei.evar_candidates }) in
let new_solution = Evd.raw_map_undefined apply_subst_einfo pruned_solution in
let new_size = Evd.fold (fun _ _ i -> i+1) new_solution 0 in
Feedback.msg_info (Pp.str (Printf.sprintf "Evars: %d -> %d\n" size new_size));
@@ -56,7 +57,7 @@ let compact el ({ solution } as pv) =
type telescope =
| TNil of Evd.evar_map
- | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+ | TCons of Environ.env * Evd.evar_map * EConstr.types * (Evd.evar_map -> EConstr.constr -> telescope)
let typeclass_resolvable = Evd.Store.field ()
@@ -72,9 +73,9 @@ let dependent_init =
| TCons (env, sigma, typ, t) ->
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (econstr, sigma, _) = Evarutil.new_evar env sigma ~src ~store typ in
+ let (gl, _) = EConstr.destEvar (Sigma.to_evar_map sigma) econstr in
let sigma = Sigma.to_evar_map sigma in
let ret, { solution = sol; comb = comb } = aux (t sigma econstr) in
- let (gl, _) = Term.destEvar econstr in
let entry = (econstr, typ) :: ret in
entry, { solution = sol; comb = gl :: comb; shelf = [] }
in
@@ -1010,7 +1011,7 @@ module Goal = struct
type ('a, 'r) t = {
env : Environ.env;
sigma : Evd.evar_map;
- concl : Term.constr ;
+ concl : EConstr.constr ;
self : Evar.t ; (* for compatibility with old-style definitions *)
}
@@ -1021,17 +1022,14 @@ module Goal = struct
let env { env=env } = env
let sigma { sigma=sigma } = Sigma.Unsafe.of_evar_map sigma
- let hyps { env=env } = Environ.named_context env
+ let hyps { env=env } = EConstr.named_context env
let concl { concl=concl } = concl
let extra { sigma=sigma; self=self } = goal_extra sigma self
- let raw_concl { concl=concl } = concl
-
-
let gmake_with info env sigma goal =
{ env = Environ.reset_with_named_context (Evd.evar_filtered_hyps info) env ;
sigma = sigma ;
- concl = Evd.evar_concl info ;
+ concl = EConstr.of_constr (Evd.evar_concl info);
self = goal }
let nf_gmake env sigma goal =
@@ -1244,12 +1242,12 @@ module V82 = struct
{ Evd.it = comb ; sigma = solution }
let top_goals initial { solution=solution; } =
- let goals = CList.map (fun (t,_) -> fst (Term.destEvar t)) initial in
+ let goals = CList.map (fun (t,_) -> fst (Term.destEvar (EConstr.Unsafe.to_constr t))) initial in
{ Evd.it = goals ; sigma=solution; }
let top_evars initial =
let evars_of_initial (c,_) =
- Evar.Set.elements (Evd.evars_of_term c)
+ Evar.Set.elements (Evd.evars_of_term (EConstr.Unsafe.to_constr c))
in
CList.flatten (CList.map evars_of_initial initial)
diff --git a/engine/proofview.mli b/engine/proofview.mli
index 294b03dca..a3b0304b1 100644
--- a/engine/proofview.mli
+++ b/engine/proofview.mli
@@ -14,6 +14,7 @@
open Util
open Term
+open EConstr
(** Main state of tactics *)
type proofview
@@ -43,7 +44,7 @@ val compact : entry -> proofview -> entry * proofview
empty [evar_map] (indeed a proof can be triggered by an incomplete
pretyping), [init] takes an additional argument to represent the
initial [evar_map]. *)
-val init : Evd.evar_map -> (Environ.env * Term.types) list -> entry * proofview
+val init : Evd.evar_map -> (Environ.env * types) list -> entry * proofview
(** A [telescope] is a list of environment and conclusion like in
{!init}, except that each element may depend on the previous
@@ -52,7 +53,7 @@ val init : Evd.evar_map -> (Environ.env * Term.types) list -> entry * proofview
[evar_map] is threaded in state passing style. *)
type telescope =
| TNil of Evd.evar_map
- | TCons of Environ.env * Evd.evar_map * Term.types * (Evd.evar_map -> Term.constr -> telescope)
+ | TCons of Environ.env * Evd.evar_map * types * (Evd.evar_map -> constr -> telescope)
(** Like {!init}, but goals are allowed to be dependent on one
another. Dependencies between goals is represented with the type
@@ -484,16 +485,12 @@ module Goal : sig
respectively the conclusion of [gl], the hypotheses of [gl], the
environment of [gl] (i.e. the global environment and the
hypotheses) and the current evar map. *)
- val concl : ([ `NF ], 'r) t -> Term.constr
- val hyps : ([ `NF ], 'r) t -> Context.Named.t
+ val concl : ('a, 'r) t -> constr
+ val hyps : ('a, 'r) t -> named_context
val env : ('a, 'r) t -> Environ.env
val sigma : ('a, 'r) t -> 'r Sigma.t
val extra : ('a, 'r) t -> Evd.Store.t
- (** Returns the goal's conclusion even if the goal is not
- normalised. *)
- val raw_concl : ('a, 'r) t -> Term.constr
-
type ('a, 'b) enter =
{ enter : 'r. ('a, 'r) t -> 'b }
diff --git a/engine/sigma.ml b/engine/sigma.ml
index 9381a33af..001f8be80 100644
--- a/engine/sigma.ml
+++ b/engine/sigma.ml
@@ -29,8 +29,8 @@ let here x s = Sigma (x, s, ())
type 'r fresh = Fresh : 's evar * 's t * ('r, 's) le -> 'r fresh
-let new_evar sigma ?naming info =
- let (sigma, evk) = Evd.new_evar sigma ?naming info in
+let new_evar sigma ?name info =
+ let (sigma, evk) = Evd.new_evar sigma ?name info in
Fresh (evk, sigma, ())
let define evk c sigma =
diff --git a/engine/sigma.mli b/engine/sigma.mli
index 7473a251b..8e8df02f2 100644
--- a/engine/sigma.mli
+++ b/engine/sigma.mli
@@ -61,7 +61,7 @@ val to_evar : 'r evar -> Evar.t
type 'r fresh = Fresh : 's evar * 's t * ('r, 's) le -> 'r fresh
-val new_evar : 'r t -> ?naming:Misctypes.intro_pattern_naming_expr ->
+val new_evar : 'r t -> ?name:Id.t ->
Evd.evar_info -> 'r fresh
val define : 'r evar -> Constr.t -> 'r t -> (unit, 'r) sigma
diff --git a/engine/termops.ml b/engine/termops.ml
index 2f4c5e204..64f4c6dc5 100644
--- a/engine/termops.ml
+++ b/engine/termops.ml
@@ -97,11 +97,334 @@ let rec pr_constr c = match kind_of_term c with
cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
str"}")
-let term_printer = ref (fun _ -> pr_constr)
-let print_constr_env t = !term_printer t
-let print_constr t = !term_printer (Global.env()) t
+let term_printer = ref (fun _env _sigma c -> pr_constr (EConstr.Unsafe.to_constr c))
+let print_constr_env env sigma t = !term_printer env sigma t
+let print_constr t = !term_printer (Global.env()) Evd.empty t
let set_print_constr f = term_printer := f
-let () = Hook.set Evd.print_constr_hook (fun env c -> !term_printer env c)
+
+module EvMap = Evar.Map
+
+let pr_evar_suggested_name evk sigma =
+ let open Evd in
+ let base_id evk' evi =
+ match evar_ident evk' sigma with
+ | Some id -> id
+ | None -> match evi.evar_source with
+ | _,Evar_kinds.ImplicitArg (c,(n,Some id),b) -> id
+ | _,Evar_kinds.VarInstance id -> id
+ | _,Evar_kinds.GoalEvar -> Id.of_string "Goal"
+ | _ ->
+ let env = reset_with_named_context evi.evar_hyps (Global.env()) in
+ Namegen.id_of_name_using_hdchar env sigma (EConstr.of_constr evi.evar_concl) Anonymous
+ in
+ let names = EvMap.mapi base_id (undefined_map sigma) in
+ let id = EvMap.find evk names in
+ let fold evk' id' (seen, n) =
+ if seen then (seen, n)
+ else if Evar.equal evk evk' then (true, n)
+ else if Id.equal id id' then (seen, succ n)
+ else (seen, n)
+ in
+ let (_, n) = EvMap.fold fold names (false, 0) in
+ if n = 0 then id else Nameops.add_suffix id (string_of_int (pred n))
+
+let pr_existential_key sigma evk =
+let open Evd in
+match evar_ident evk sigma with
+| None ->
+ str "?" ++ pr_id (pr_evar_suggested_name evk sigma)
+| Some id ->
+ str "?" ++ pr_id id
+
+let pr_instance_status (sc,typ) =
+ let open Evd in
+ begin match sc with
+ | IsSubType -> str " [or a subtype of it]"
+ | IsSuperType -> str " [or a supertype of it]"
+ | Conv -> mt ()
+ end ++
+ begin match typ with
+ | CoerceToType -> str " [up to coercion]"
+ | TypeNotProcessed -> mt ()
+ | TypeProcessed -> str " [type is checked]"
+ end
+
+let protect f x =
+ try f x
+ with e -> str "EXCEPTION: " ++ str (Printexc.to_string e)
+
+let print_kconstr a =
+ protect (fun c -> print_constr (EConstr.of_constr c)) a
+
+let pr_meta_map evd =
+ let open Evd in
+ let print_constr = print_kconstr in
+ let pr_name = function
+ Name id -> str"[" ++ pr_id id ++ str"]"
+ | _ -> mt() in
+ let pr_meta_binding = function
+ | (mv,Cltyp (na,b)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " : " ++
+ print_constr b.rebus ++ fnl ())
+ | (mv,Clval(na,(b,s),t)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " := " ++
+ print_constr b.rebus ++
+ str " : " ++ print_constr t.rebus ++
+ spc () ++ pr_instance_status s ++ fnl ())
+ in
+ prlist pr_meta_binding (meta_list evd)
+
+let pr_decl (decl,ok) =
+ let open NamedDecl in
+ let print_constr = print_kconstr in
+ match decl with
+ | LocalAssum (id,_) -> if ok then pr_id id else (str "{" ++ pr_id id ++ str "}")
+ | LocalDef (id,c,_) -> str (if ok then "(" else "{") ++ pr_id id ++ str ":=" ++
+ print_constr c ++ str (if ok then ")" else "}")
+
+let pr_evar_source = function
+ | Evar_kinds.NamedHole id -> pr_id id
+ | Evar_kinds.QuestionMark _ -> str "underscore"
+ | Evar_kinds.CasesType false -> str "pattern-matching return predicate"
+ | Evar_kinds.CasesType true ->
+ str "subterm of pattern-matching return predicate"
+ | Evar_kinds.BinderType (Name id) -> str "type of " ++ Nameops.pr_id id
+ | Evar_kinds.BinderType Anonymous -> str "type of anonymous binder"
+ | Evar_kinds.ImplicitArg (c,(n,ido),b) ->
+ let open Globnames in
+ let print_constr = print_kconstr in
+ let id = Option.get ido in
+ str "parameter " ++ pr_id id ++ spc () ++ str "of" ++
+ spc () ++ print_constr (printable_constr_of_global c)
+ | Evar_kinds.InternalHole -> str "internal placeholder"
+ | Evar_kinds.TomatchTypeParameter (ind,n) ->
+ let print_constr = print_kconstr in
+ pr_nth n ++ str " argument of type " ++ print_constr (mkInd ind)
+ | Evar_kinds.GoalEvar -> str "goal evar"
+ | Evar_kinds.ImpossibleCase -> str "type of impossible pattern-matching clause"
+ | Evar_kinds.MatchingVar _ -> str "matching variable"
+ | Evar_kinds.VarInstance id -> str "instance of " ++ pr_id id
+ | Evar_kinds.SubEvar evk ->
+ let open Evd in
+ str "subterm of " ++ str (string_of_existential evk)
+
+let pr_evar_info evi =
+ let open Evd in
+ let print_constr = print_kconstr in
+ let phyps =
+ try
+ let decls = match Filter.repr (evar_filter evi) with
+ | None -> List.map (fun c -> (c, true)) (evar_context evi)
+ | Some filter -> List.combine (evar_context evi) filter
+ in
+ prlist_with_sep spc pr_decl (List.rev decls)
+ with Invalid_argument _ -> str "Ill-formed filtered context" in
+ let pty = print_constr evi.evar_concl in
+ let pb =
+ match evi.evar_body with
+ | Evar_empty -> mt ()
+ | Evar_defined c -> spc() ++ str"=> " ++ print_constr c
+ in
+ let candidates =
+ match evi.evar_body, evi.evar_candidates with
+ | Evar_empty, Some l ->
+ spc () ++ str "{" ++
+ prlist_with_sep (fun () -> str "|") print_constr l ++ str "}"
+ | _ ->
+ mt ()
+ in
+ let src = str "(" ++ pr_evar_source (snd evi.evar_source) ++ str ")" in
+ hov 2
+ (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]" ++
+ candidates ++ spc() ++ src)
+
+let compute_evar_dependency_graph sigma =
+ let open Evd in
+ (* Compute the map binding ev to the evars whose body depends on ev *)
+ let fold evk evi acc =
+ let fold_ev evk' acc =
+ let tab =
+ try EvMap.find evk' acc
+ with Not_found -> Evar.Set.empty
+ in
+ EvMap.add evk' (Evar.Set.add evk tab) acc
+ in
+ match evar_body evi with
+ | Evar_empty -> acc
+ | Evar_defined c -> Evar.Set.fold fold_ev (evars_of_term c) acc
+ in
+ Evd.fold fold sigma EvMap.empty
+
+let evar_dependency_closure n sigma =
+ let open Evd in
+ (** Create the DAG of depth [n] representing the recursive dependencies of
+ undefined evars. *)
+ let graph = compute_evar_dependency_graph sigma in
+ let rec aux n curr accu =
+ if Int.equal n 0 then Evar.Set.union curr accu
+ else
+ let fold evk accu =
+ try
+ let deps = EvMap.find evk graph in
+ Evar.Set.union deps accu
+ with Not_found -> accu
+ in
+ (** Consider only the newly added evars *)
+ let ncurr = Evar.Set.fold fold curr Evar.Set.empty in
+ (** Merge the others *)
+ let accu = Evar.Set.union curr accu in
+ aux (n - 1) ncurr accu
+ in
+ let undef = EvMap.domain (undefined_map sigma) in
+ aux n undef Evar.Set.empty
+
+let evar_dependency_closure n sigma =
+ let open Evd in
+ let deps = evar_dependency_closure n sigma in
+ let map = EvMap.bind (fun ev -> find sigma ev) deps in
+ EvMap.bindings map
+
+let has_no_evar sigma =
+ try let () = Evd.fold (fun _ _ () -> raise Exit) sigma () in true
+ with Exit -> false
+
+let pr_evd_level evd = UState.pr_uctx_level (Evd.evar_universe_context evd)
+
+let pr_evar_universe_context ctx =
+ let open UState in
+ let open Evd in
+ let prl = pr_uctx_level ctx in
+ if UState.is_empty ctx then mt ()
+ else
+ (str"UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.pr_universe_context_set prl (evar_universe_context_set ctx)) ++ fnl () ++
+ str"ALGEBRAIC UNIVERSES:"++brk(0,1)++
+ h 0 (Univ.LSet.pr prl (UState.algebraics ctx)) ++ fnl() ++
+ str"UNDEFINED UNIVERSES:"++brk(0,1)++
+ h 0 (Universes.pr_universe_opt_subst (UState.subst ctx)) ++ fnl())
+
+let print_env_short env =
+ let print_constr = print_kconstr in
+ let pr_rel_decl = function
+ | RelDecl.LocalAssum (n,_) -> pr_name n
+ | RelDecl.LocalDef (n,b,_) -> str "(" ++ pr_name n ++ str " := " ++ print_constr b ++ str ")"
+ in
+ let pr_named_decl = NamedDecl.to_rel_decl %> pr_rel_decl in
+ let nc = List.rev (named_context env) in
+ let rc = List.rev (rel_context env) in
+ str "[" ++ pr_sequence pr_named_decl nc ++ str "]" ++ spc () ++
+ str "[" ++ pr_sequence pr_rel_decl rc ++ str "]"
+
+let pr_evar_constraints sigma pbs =
+ let pr_evconstr (pbty, env, t1, t2) =
+ let env =
+ (** We currently allow evar instances to refer to anonymous de
+ Bruijn indices, so we protect the error printing code in this
+ case by giving names to every de Bruijn variable in the
+ rel_context of the conversion problem. MS: we should rather
+ stop depending on anonymous variables, they can be used to
+ indicate independency. Also, this depends on a strategy for
+ naming/renaming. *)
+ Namegen.make_all_name_different env sigma
+ in
+ print_env_short env ++ spc () ++ str "|-" ++ spc () ++
+ print_constr_env env sigma (EConstr.of_constr t1) ++ spc () ++
+ str (match pbty with
+ | Reduction.CONV -> "=="
+ | Reduction.CUMUL -> "<=") ++
+ spc () ++ print_constr_env env Evd.empty (EConstr.of_constr t2)
+ in
+ prlist_with_sep fnl pr_evconstr pbs
+
+let pr_evar_map_gen with_univs pr_evars sigma =
+ let uvs = Evd.evar_universe_context sigma in
+ let (_, conv_pbs) = Evd.extract_all_conv_pbs sigma in
+ let evs = if has_no_evar sigma then mt () else pr_evars sigma ++ fnl ()
+ and svs = if with_univs then pr_evar_universe_context uvs else mt ()
+ and cstrs =
+ if List.is_empty conv_pbs then mt ()
+ else
+ str "CONSTRAINTS:" ++ brk (0, 1) ++
+ pr_evar_constraints sigma conv_pbs ++ fnl ()
+ and metas =
+ if List.is_empty (Evd.meta_list sigma) then mt ()
+ else
+ str "METAS:" ++ brk (0, 1) ++ pr_meta_map sigma
+ in
+ evs ++ svs ++ cstrs ++ metas
+
+let pr_evar_list sigma l =
+ let open Evd in
+ let pr (ev, evi) =
+ h 0 (str (string_of_existential ev) ++
+ str "==" ++ pr_evar_info evi ++
+ (if evi.evar_body == Evar_empty
+ then str " {" ++ pr_existential_key sigma ev ++ str "}"
+ else mt ()))
+ in
+ h 0 (prlist_with_sep fnl pr l)
+
+let pr_evar_by_depth depth sigma = match depth with
+| None ->
+ (* Print all evars *)
+ let to_list d =
+ let open Evd in
+ (* Workaround for change in Map.fold behavior in ocaml 3.08.4 *)
+ let l = ref [] in
+ let fold_def evk evi () = match evi.evar_body with
+ | Evar_defined _ -> l := (evk, evi) :: !l
+ | Evar_empty -> ()
+ in
+ let fold_undef evk evi () = match evi.evar_body with
+ | Evar_empty -> l := (evk, evi) :: !l
+ | Evar_defined _ -> ()
+ in
+ Evd.fold fold_def d ();
+ Evd.fold fold_undef d ();
+ !l
+ in
+ str"EVARS:"++brk(0,1)++pr_evar_list sigma (to_list sigma)++fnl()
+| Some n ->
+ (* Print all evars *)
+ str"UNDEFINED EVARS:"++
+ (if Int.equal n 0 then mt() else str" (+level "++int n++str" closure):")++
+ brk(0,1)++
+ pr_evar_list sigma (evar_dependency_closure n sigma)++fnl()
+
+let pr_evar_by_filter filter sigma =
+ let open Evd in
+ let elts = Evd.fold (fun evk evi accu -> (evk, evi) :: accu) sigma [] in
+ let elts = List.rev elts in
+ let is_def (_, evi) = match evi.evar_body with
+ | Evar_defined _ -> true
+ | Evar_empty -> false
+ in
+ let (defined, undefined) = List.partition is_def elts in
+ let filter (evk, evi) = filter evk evi in
+ let defined = List.filter filter defined in
+ let undefined = List.filter filter undefined in
+ let prdef =
+ if List.is_empty defined then mt ()
+ else str "DEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma defined
+ in
+ let prundef =
+ if List.is_empty undefined then mt ()
+ else str "UNDEFINED EVARS:" ++ brk (0, 1) ++
+ pr_evar_list sigma undefined
+ in
+ prdef ++ prundef
+
+let pr_evar_map ?(with_univs=true) depth sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_depth depth sigma) sigma
+
+let pr_evar_map_filter ?(with_univs=true) filter sigma =
+ pr_evar_map_gen with_univs (fun sigma -> pr_evar_by_filter filter sigma) sigma
+
+let pr_metaset metas =
+ str "[" ++ pr_sequence pr_meta (Evd.Metaset.elements metas) ++ str "]"
let pr_var_decl env decl =
let open NamedDecl in
@@ -109,9 +432,10 @@ let pr_var_decl env decl =
| LocalAssum _ -> mt ()
| LocalDef (_,c,_) ->
(* Force evaluation *)
- let pb = print_constr_env env c in
+ let c = EConstr.of_constr c in
+ let pb = print_constr_env env Evd.empty c in
(str" := " ++ pb ++ cut () ) in
- let pt = print_constr_env env (get_type decl) in
+ let pt = print_constr_env env Evd.empty (EConstr.of_constr (get_type decl)) in
let ptyp = (str" : " ++ pt) in
(pr_id (get_id decl) ++ hov 0 (pbody ++ ptyp))
@@ -121,9 +445,10 @@ let pr_rel_decl env decl =
| LocalAssum _ -> mt ()
| LocalDef (_,c,_) ->
(* Force evaluation *)
- let pb = print_constr_env env c in
+ let c = EConstr.of_constr c in
+ let pb = print_constr_env env Evd.empty c in
(str":=" ++ spc () ++ pb ++ spc ()) in
- let ptyp = print_constr_env env (get_type decl) in
+ let ptyp = print_constr_env env Evd.empty (EConstr.of_constr (get_type decl)) in
match get_name decl with
| Anonymous -> hov 0 (str"<>" ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
| Name id -> hov 0 (pr_id id ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
@@ -159,6 +484,7 @@ let print_env env =
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i))
let rel_list n m =
+ let open EConstr in
let rec reln l p =
if p>m then l else reln (mkRel(n+p)::l) (p+1)
in
@@ -166,6 +492,7 @@ let rel_list n m =
let push_rel_assum (x,t) env =
let open RelDecl in
+ let open EConstr in
push_rel (LocalAssum (x,t)) env
let push_rels_assum assums =
@@ -199,21 +526,17 @@ let lookup_rel_id id sign =
lookrec 1 sign
(* Constructs either [forall x:t, c] or [let x:=b:t in c] *)
-let mkProd_or_LetIn decl c =
- let open RelDecl in
- match decl with
- | LocalAssum (na,t) -> mkProd (na, t, c)
- | LocalDef (na,b,t) -> mkLetIn (na, b, t, c)
-
+let mkProd_or_LetIn = EConstr.mkProd_or_LetIn
(* Constructs either [forall x:t, c] or [c] in which [x] is replaced by [b] *)
let mkProd_wo_LetIn decl c =
+ let open EConstr in
let open RelDecl in
match decl with
| LocalAssum (na,t) -> mkProd (na, t, c)
- | LocalDef (_,b,_) -> subst1 b c
+ | LocalDef (_,b,_) -> Vars.subst1 b c
-let it_mkProd init = List.fold_left (fun c (n,t) -> mkProd (n, t, c)) init
-let it_mkLambda init = List.fold_left (fun c (n,t) -> mkLambda (n, t, c)) init
+let it_mkProd init = List.fold_left (fun c (n,t) -> EConstr.mkProd (n, t, c)) init
+let it_mkLambda init = List.fold_left (fun c (n,t) -> EConstr.mkLambda (n, t, c)) init
let it_named_context_quantifier f ~init =
List.fold_left (fun c d -> f d c) init
@@ -221,9 +544,9 @@ let it_named_context_quantifier f ~init =
let it_mkProd_or_LetIn init = it_named_context_quantifier mkProd_or_LetIn ~init
let it_mkProd_wo_LetIn init = it_named_context_quantifier mkProd_wo_LetIn ~init
let it_mkLambda_or_LetIn init = it_named_context_quantifier mkLambda_or_LetIn ~init
-let it_mkNamedProd_or_LetIn init = it_named_context_quantifier mkNamedProd_or_LetIn ~init
+let it_mkNamedProd_or_LetIn init = it_named_context_quantifier EConstr.mkNamedProd_or_LetIn ~init
let it_mkNamedProd_wo_LetIn init = it_named_context_quantifier mkNamedProd_wo_LetIn ~init
-let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier mkNamedLambda_or_LetIn ~init
+let it_mkNamedLambda_or_LetIn init = it_named_context_quantifier EConstr.mkNamedLambda_or_LetIn ~init
let it_mkLambda_or_LetIn_from_no_LetIn c decls =
let open RelDecl in
@@ -236,37 +559,39 @@ let it_mkLambda_or_LetIn_from_no_LetIn c decls =
(* *)
(* strips head casts and flattens head applications *)
-let rec strip_head_cast c = match kind_of_term c with
+let rec strip_head_cast sigma c = match EConstr.kind sigma c with
| App (f,cl) ->
- let rec collapse_rec f cl2 = match kind_of_term f with
+ let rec collapse_rec f cl2 = match EConstr.kind sigma f with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
| Cast (c,_,_) -> collapse_rec c cl2
- | _ -> if Int.equal (Array.length cl2) 0 then f else mkApp (f,cl2)
+ | _ -> if Int.equal (Array.length cl2) 0 then f else EConstr.mkApp (f,cl2)
in
collapse_rec f cl
- | Cast (c,_,_) -> strip_head_cast c
+ | Cast (c,_,_) -> strip_head_cast sigma c
| _ -> c
-let rec drop_extra_implicit_args c = match kind_of_term c with
+let rec drop_extra_implicit_args sigma c = match EConstr.kind sigma c with
(* Removed trailing extra implicit arguments, what improves compatibility
for constants with recently added maximal implicit arguments *)
- | App (f,args) when isEvar (Array.last args) ->
- drop_extra_implicit_args
+ | App (f,args) when EConstr.isEvar sigma (Array.last args) ->
+ let open EConstr in
+ drop_extra_implicit_args sigma
(mkApp (f,fst (Array.chop (Array.length args - 1) args)))
| _ -> c
(* Get the last arg of an application *)
-let last_arg c = match kind_of_term c with
+let last_arg sigma c = match EConstr.kind sigma c with
| App (f,cl) -> Array.last cl
| _ -> anomaly (Pp.str "last_arg")
(* Get the last arg of an application *)
-let decompose_app_vect c =
- match kind_of_term c with
+let decompose_app_vect sigma c =
+ match EConstr.kind sigma c with
| App (f,cl) -> (f, cl)
| _ -> (c,[||])
let adjust_app_list_size f1 l1 f2 l2 =
+ let open EConstr in
let len1 = List.length l1 and len2 = List.length l2 in
if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
@@ -277,14 +602,15 @@ let adjust_app_list_size f1 l1 f2 l2 =
(applist (f1,extras), restl1, f2, l2)
let adjust_app_array_size f1 l1 f2 l2 =
+ let open EConstr in
let len1 = Array.length l1 and len2 = Array.length l2 in
if Int.equal len1 len2 then (f1,l1,f2,l2)
else if len1 < len2 then
let extras,restl2 = Array.chop (len2-len1) l2 in
- (f1, l1, appvect (f2,extras), restl2)
+ (f1, l1, mkApp (f2,extras), restl2)
else
let extras,restl1 = Array.chop (len1-len2) l1 in
- (appvect (f1,extras), restl1, f2, l2)
+ (mkApp (f1,extras), restl1, f2, l2)
(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
subterms of [c]; it carries an extra data [l] (typically a name
@@ -321,6 +647,8 @@ let map_constr_with_named_binders g f l c = match kind_of_term c with
(co-)fixpoint) *)
let fold_rec_types g (lna,typarray,_) e =
+ let open EConstr in
+ let open Vars in
let ctxt = Array.map2_i (fun i na t -> RelDecl.LocalAssum (na, lift i t)) lna typarray in
Array.fold_left (fun e assum -> g assum e) e ctxt
@@ -336,9 +664,10 @@ let map_left2 f a g b =
r, s
end
-let map_constr_with_binders_left_to_right g f l c =
+let map_constr_with_binders_left_to_right sigma g f l c =
let open RelDecl in
- match kind_of_term c with
+ let open EConstr in
+ match EConstr.kind sigma c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
| Cast (b,k,t) ->
@@ -400,9 +729,10 @@ let map_constr_with_binders_left_to_right g f l c =
else mkCoFix (ln,(lna,tl',bl'))
(* strong *)
-let map_constr_with_full_binders g f l cstr =
+let map_constr_with_full_binders sigma g f l cstr =
+ let open EConstr in
let open RelDecl in
- match kind_of_term cstr with
+ match EConstr.kind sigma cstr with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> cstr
| Cast (c,k, t) ->
@@ -411,16 +741,16 @@ let map_constr_with_full_binders g f l cstr =
if c==c' && t==t' then cstr else mkCast (c', k, t')
| Prod (na,t,c) ->
let t' = f l t in
- let c' = f (g (LocalAssum (na,t)) l) c in
+ let c' = f (g (LocalAssum (na, t)) l) c in
if t==t' && c==c' then cstr else mkProd (na, t', c')
| Lambda (na,t,c) ->
let t' = f l t in
- let c' = f (g (LocalAssum (na,t)) l) c in
+ let c' = f (g (LocalAssum (na, t)) l) c in
if t==t' && c==c' then cstr else mkLambda (na, t', c')
| LetIn (na,b,t,c) ->
let b' = f l b in
let t' = f l t in
- let c' = f (g (LocalDef (na,b,t)) l) c in
+ let c' = f (g (LocalDef (na, b, t)) l) c in
if b==b' && t==t' && c==c' then cstr else mkLetIn (na, b', t', c')
| App (c,al) ->
let c' = f l c in
@@ -441,7 +771,7 @@ let map_constr_with_full_binders g f l cstr =
| Fix (ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na, t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -449,7 +779,7 @@ let map_constr_with_full_binders g f l cstr =
| CoFix(ln,(lna,tl,bl)) ->
let tl' = Array.map (f l) tl in
let l' =
- Array.fold_left2 (fun l na t -> g (LocalAssum (na,t)) l) l lna tl in
+ Array.fold_left2 (fun l na t -> g (LocalAssum (na, t)) l) l lna tl in
let bl' = Array.map (f l') bl in
if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl'
then cstr
@@ -462,30 +792,31 @@ let map_constr_with_full_binders g f l cstr =
index) which is processed by [g] (which typically add 1 to [n]) at
each binder traversal; it is not recursive *)
-let fold_constr_with_full_binders g f n acc c =
+let fold_constr_with_full_binders sigma g f n acc c =
let open RelDecl in
- match kind_of_term c with
+ let inj c = EConstr.Unsafe.to_constr c in
+ match EConstr.kind sigma 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
- | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
- | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c
+ | Prod (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
+ | Lambda (na,t,c) -> f (g (LocalAssum (na, inj t)) n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g (LocalDef (na, inj b, inj t)) n) (f n (f n acc b) t) c
| App (c,l) -> Array.fold_left (f n) (f n acc c) l
| Proj (p,c) -> f n acc c
| Evar (_,l) -> Array.fold_left (f n) acc l
| Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
| Fix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
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
| CoFix (_,(lna,tl,bl)) ->
- let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n, inj t)) c) n lna tl in
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 fold_constr_with_binders g f n acc c =
- fold_constr_with_full_binders (fun _ x -> g x) f n acc c
+let fold_constr_with_binders sigma g f n acc c =
+ fold_constr_with_full_binders sigma (fun _ x -> g x) f n acc c
(* [iter_constr_with_full_binders g f acc c] iters [f acc] on the immediate
subterms of [c]; it carries an extra data [acc] which is processed by [g] at
@@ -520,29 +851,29 @@ let iter_constr_with_full_binders g f l c =
exception Occur
-let occur_meta c =
- let rec occrec c = match kind_of_term c with
+let occur_meta sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Meta _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_existential c =
- let rec occrec c = match kind_of_term c with
+let occur_existential sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Evar _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_meta_or_existential c =
- let rec occrec c = match kind_of_term c with
+let occur_meta_or_existential sigma c =
+ let rec occrec c = match EConstr.kind sigma c with
| Evar _ -> raise Occur
| Meta _ -> raise Occur
- | _ -> iter_constr occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
-let occur_evar n c =
- let rec occur_rec c = match kind_of_term c with
+let occur_evar sigma n c =
+ let rec occur_rec c = match EConstr.kind sigma c with
| Evar (sp,_) when Evar.equal sp n -> raise Occur
- | _ -> iter_constr occur_rec c
+ | _ -> EConstr.iter sigma occur_rec c
in
try occur_rec c; false with Occur -> true
@@ -550,55 +881,55 @@ let occur_in_global env id constr =
let vars = vars_of_global env constr in
if Id.Set.mem id vars then raise Occur
-let occur_var env id c =
+let occur_var env sigma id c =
let rec occur_rec c =
- match kind_of_term c with
- | Var _ | Const _ | Ind _ | Construct _ -> occur_in_global env id c
- | _ -> iter_constr occur_rec c
+ match EConstr.kind sigma c with
+ | Var _ | Const _ | Ind _ | Construct _ -> occur_in_global env id (EConstr.to_constr sigma c)
+ | _ -> EConstr.iter sigma occur_rec c
in
try occur_rec c; false with Occur -> true
-let occur_var_in_decl env hyp decl =
+let occur_var_in_decl env sigma hyp decl =
let open NamedDecl in
match decl with
- | LocalAssum (_,typ) -> occur_var env hyp typ
+ | LocalAssum (_,typ) -> occur_var env sigma hyp typ
| LocalDef (_, body, typ) ->
- occur_var env hyp typ ||
- occur_var env hyp body
+ occur_var env sigma hyp typ ||
+ occur_var env sigma hyp body
-let local_occur_var id c =
- let rec occur c = match kind_of_term c with
+let local_occur_var sigma id c =
+ let rec occur c = match EConstr.kind sigma c with
| Var id' -> if Id.equal id id' then raise Occur
- | _ -> Constr.iter occur c
+ | _ -> EConstr.iter sigma occur c
in
try occur c; false with Occur -> true
(* returns the list of free debruijn indices in a term *)
-let free_rels m =
- let rec frec depth acc c = match kind_of_term c with
+let free_rels sigma m =
+ let rec frec depth acc c = match EConstr.kind sigma c with
| Rel n -> if n >= depth then Int.Set.add (n-depth+1) acc else acc
- | _ -> fold_constr_with_binders succ frec depth acc c
+ | _ -> fold_constr_with_binders sigma succ frec depth acc c
in
frec 1 Int.Set.empty m
(* collects all metavar occurrences, in left-to-right order, preserving
* repetitions and all. *)
-let collect_metas c =
+let collect_metas sigma c =
let rec collrec acc c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Meta mv -> List.add_set Int.equal mv acc
- | _ -> fold_constr collrec acc c
+ | _ -> EConstr.fold sigma collrec acc c
in
List.rev (collrec [] c)
(* collects all vars; warning: this is only visible vars, not dependencies in
all section variables; for the latter, use global_vars_set *)
-let collect_vars c =
- let rec aux vars c = match kind_of_term c with
+let collect_vars sigma c =
+ let rec aux vars c = match EConstr.kind sigma c with
| Var id -> Id.Set.add id vars
- | _ -> fold_constr aux vars c in
+ | _ -> EConstr.fold sigma aux vars c in
aux Id.Set.empty c
let vars_of_global_reference env gr =
@@ -608,54 +939,56 @@ let vars_of_global_reference env gr =
(* Tests whether [m] is a subterm of [t]:
[m] is appropriately lifted through abstractions of [t] *)
-let dependent_main noevar univs m t =
+let dependent_main noevar univs sigma m t =
+ let open EConstr in
let eqc x y =
- if univs then not (Option.is_empty (Universes.eq_constr_universes x y))
- else eq_constr_nounivs x y
+ if univs then not (Option.is_empty (eq_constr_universes sigma x y))
+ else eq_constr_nounivs sigma x y
in
let rec deprec m t =
if eqc m t then
raise Occur
else
- match kind_of_term m, kind_of_term t with
+ match EConstr.kind sigma m, EConstr.kind sigma t with
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
deprec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
CArray.Fun1.iter deprec m
(Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
- | _, Cast (c,_,_) when noevar && isMeta c -> ()
+ | _, Cast (c,_,_) when noevar && isMeta sigma c -> ()
| _, Evar _ when noevar -> ()
- | _ -> iter_constr_with_binders (fun c -> lift 1 c) deprec m t
+ | _ -> EConstr.iter_with_binders sigma (fun c -> Vars.lift 1 c) deprec m t
in
try deprec m t; false with Occur -> true
-let dependent = dependent_main false false
-let dependent_no_evar = dependent_main true false
+let dependent sigma c t = dependent_main false false sigma c t
+let dependent_no_evar sigma c t = dependent_main true false sigma c t
-let dependent_univs = dependent_main false true
-let dependent_univs_no_evar = dependent_main true true
+let dependent_univs sigma c t = dependent_main false true sigma c t
+let dependent_univs_no_evar sigma c t = dependent_main true true sigma c t
-let dependent_in_decl a decl =
+let dependent_in_decl sigma a decl =
let open NamedDecl in
match decl with
- | LocalAssum (_,t) -> dependent a t
- | LocalDef (_, body, t) -> dependent a body || dependent a t
+ | LocalAssum (_,t) -> dependent sigma a t
+ | LocalDef (_, body, t) -> dependent sigma a body || dependent sigma a t
-let count_occurrences m t =
+let count_occurrences sigma m t =
+ let open EConstr in
let n = ref 0 in
let rec countrec m t =
- if eq_constr m t then
+ if EConstr.eq_constr sigma m t then
incr n
else
- match kind_of_term m, kind_of_term t with
+ match EConstr.kind sigma m, EConstr.kind sigma t with
| App (fm,lm), App (ft,lt) when Array.length lm < Array.length lt ->
countrec m (mkApp (ft,Array.sub lt 0 (Array.length lm)));
Array.iter (countrec m)
(Array.sub lt
(Array.length lm) ((Array.length lt) - (Array.length lm)))
- | _, Cast (c,_,_) when isMeta c -> ()
+ | _, Cast (c,_,_) when isMeta sigma c -> ()
| _, Evar _ -> ()
- | _ -> iter_constr_with_binders (lift 1) countrec m t
+ | _ -> EConstr.iter_with_binders sigma (Vars.lift 1) countrec m t
in
countrec m t;
!n
@@ -663,7 +996,7 @@ let count_occurrences m t =
(* Synonymous *)
let occur_term = dependent
-let pop t = lift (-1) t
+let pop t = EConstr.Vars.lift (-1) t
(***************************)
(* bindings functions *)
@@ -673,50 +1006,55 @@ type meta_type_map = (metavariable * types) list
type meta_value_map = (metavariable * constr) list
+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
| Meta i -> (try Int.List.assoc i bl with Not_found -> c)
| _ -> map_constr (subst_meta bl) c
-let rec strip_outer_cast c = match kind_of_term c with
- | Cast (c,_,_) -> strip_outer_cast c
+let rec strip_outer_cast sigma c = match EConstr.kind sigma c with
+ | Cast (c,_,_) -> strip_outer_cast sigma c
| _ -> c
(* flattens application lists throwing casts in-between *)
-let collapse_appl c = match kind_of_term c with
+let collapse_appl sigma c = match EConstr.kind sigma c with
| App (f,cl) ->
let rec collapse_rec f cl2 =
- match kind_of_term (strip_outer_cast f) with
+ match EConstr.kind sigma (strip_outer_cast sigma f) with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
- | _ -> mkApp (f,cl2)
+ | _ -> EConstr.mkApp (f,cl2)
in
collapse_rec f cl
| _ -> c
(* First utilities for avoiding telescope computation for subst_term *)
-let prefix_application eq_fun (k,c) (t : constr) =
- let c' = collapse_appl c and t' = collapse_appl t in
- match kind_of_term c', kind_of_term t' with
+let prefix_application sigma eq_fun (k,c) t =
+ let open EConstr in
+ let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
+ match EConstr.kind sigma c', EConstr.kind sigma t' with
| App (f1,cl1), App (f2,cl2) ->
let l1 = Array.length cl1
and l2 = Array.length cl2 in
if l1 <= l2
- && eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
+ && eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
else
None
| _ -> None
-let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
- let c' = collapse_appl c and t' = collapse_appl t in
- match kind_of_term c', kind_of_term t' with
+let my_prefix_application sigma eq_fun (k,c) by_c t =
+ let open EConstr in
+ let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
+ match EConstr.kind sigma c', EConstr.kind sigma t' with
| App (f1,cl1), App (f2,cl2) ->
let l1 = Array.length cl1
and l2 = Array.length cl2 in
if l1 <= l2
- && eq_fun c' (mkApp (f2, Array.sub cl2 0 l1)) then
- Some (mkApp ((lift k by_c), Array.sub cl2 l1 (l2 - l1)))
+ && eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
+ Some (mkApp ((Vars.lift k by_c), Array.sub cl2 l1 (l2 - l1)))
else
None
| _ -> None
@@ -725,35 +1063,37 @@ let my_prefix_application eq_fun (k,c) (by_c : constr) (t : constr) =
substitutes [(Rel 1)] for all occurrences of term [c] in a term [t];
works if [c] has rels *)
-let subst_term_gen eq_fun c t =
+let subst_term_gen sigma eq_fun c t =
+ let open EConstr in
+ let open Vars in
let rec substrec (k,c as kc) t =
- match prefix_application eq_fun kc t with
+ match prefix_application sigma eq_fun kc t with
| Some x -> x
| None ->
- if eq_fun c t then mkRel k
+ if eq_fun sigma c t then mkRel k
else
- map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
+ EConstr.map_with_binders sigma (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
in
substrec (1,c) t
-let subst_term = subst_term_gen eq_constr
+let subst_term sigma c t = subst_term_gen sigma EConstr.eq_constr c t
(* Recognizing occurrences of a given subterm in a term :
[replace_term c1 c2 t] substitutes [c2] for all occurrences of
term [c1] in a term [t]; works if [c1] and [c2] have rels *)
-let replace_term_gen eq_fun c by_c in_t =
+let replace_term_gen sigma eq_fun c by_c in_t =
let rec substrec (k,c as kc) t =
- match my_prefix_application eq_fun kc by_c t with
+ match my_prefix_application sigma eq_fun kc by_c t with
| Some x -> x
| None ->
- (if eq_fun c t then (lift k by_c) else
- map_constr_with_binders (fun (k,c) -> (k+1,lift 1 c))
+ (if eq_fun sigma c t then (EConstr.Vars.lift k by_c) else
+ EConstr.map_with_binders sigma (fun (k,c) -> (k+1,EConstr.Vars.lift 1 c))
substrec kc t)
in
substrec (0,c) in_t
-let replace_term = replace_term_gen eq_constr
+let replace_term sigma c byc t = replace_term_gen sigma EConstr.eq_constr c byc t
let vars_of_env env =
let s =
@@ -804,15 +1144,37 @@ let is_section_variable id =
try let _ = Global.lookup_named id in true
with Not_found -> false
-let isGlobalRef c =
- match kind_of_term c with
+let global_of_constr sigma c =
+ let open Globnames in
+ match EConstr.kind sigma c with
+ | Const (c, u) -> ConstRef c, u
+ | Ind (i, u) -> IndRef i, u
+ | Construct (c, u) -> ConstructRef c, u
+ | Var id -> VarRef id, EConstr.EInstance.empty
+ | _ -> raise Not_found
+
+let is_global sigma c t =
+ let open Globnames in
+ match c, EConstr.kind sigma 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'
+ | VarRef id, Var id' -> Id.equal id id'
+ | _ -> false
+
+let isGlobalRef sigma c =
+ match EConstr.kind sigma c with
| Const _ | Ind _ | Construct _ | Var _ -> true
| _ -> false
-let is_template_polymorphic env f =
- match kind_of_term f with
- | Ind ind -> Environ.template_polymorphic_pind ind env
- | Const c -> Environ.template_polymorphic_pconstant c env
+let is_template_polymorphic env sigma f =
+ match EConstr.kind sigma f with
+ | Ind (ind, u) ->
+ if not (EConstr.EInstance.is_empty u) then false
+ else Environ.template_polymorphic_ind ind env
+ | Const (cst, u) ->
+ if not (EConstr.EInstance.is_empty u) then false
+ else Environ.template_polymorphic_constant cst env
| _ -> false
let base_sort_cmp pb s0 s1 =
@@ -822,21 +1184,33 @@ let base_sort_cmp pb s0 s1 =
| (Type u1, Type u2) -> true
| _ -> false
+let rec is_Prop sigma c = match EConstr.kind sigma c with
+ | Sort u ->
+ begin match EConstr.ESorts.kind sigma u with
+ | Prop Null -> true
+ | _ -> false
+ end
+ | Cast (c,_,_) -> is_Prop sigma c
+ | _ -> false
+
(* eq_constr extended with universe erasure *)
-let compare_constr_univ f cv_pb t1 t2 =
- match kind_of_term t1, kind_of_term t2 with
- Sort s1, Sort s2 -> base_sort_cmp cv_pb s1 s2
+let compare_constr_univ sigma f cv_pb t1 t2 =
+ let open EConstr in
+ match EConstr.kind sigma t1, EConstr.kind sigma t2 with
+ Sort s1, Sort s2 -> base_sort_cmp cv_pb (ESorts.kind sigma s1) (ESorts.kind sigma s2)
| Prod (_,t1,c1), Prod (_,t2,c2) ->
f Reduction.CONV t1 t2 && f cv_pb c1 c2
- | _ -> compare_constr (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2
+ | _ -> EConstr.compare_constr sigma (fun t1 t2 -> f Reduction.CONV t1 t2) t1 t2
-let rec constr_cmp cv_pb t1 t2 = compare_constr_univ constr_cmp cv_pb t1 t2
+let constr_cmp sigma cv_pb t1 t2 =
+ let rec compare cv_pb t1 t2 = compare_constr_univ sigma compare cv_pb t1 t2 in
+ compare cv_pb t1 t2
-let eq_constr t1 t2 = constr_cmp Reduction.CONV t1 t2
+let eq_constr sigma t1 t2 = constr_cmp sigma Reduction.CONV t1 t2
(* App(c,[t1,...tn]) -> ([c,t1,...,tn-1],tn)
App(c,[||]) -> ([],c) *)
-let split_app c = match kind_of_term c with
+let split_app sigma c = match EConstr.kind sigma c with
App(c,l) ->
let len = Array.length l in
if Int.equal len 0 then ([],c) else
@@ -845,28 +1219,30 @@ let split_app c = match kind_of_term c with
c::(Array.to_list prev), last
| _ -> assert false
-type subst = (Context.Rel.t * constr) Evar.Map.t
+type subst = (EConstr.rel_context * EConstr.constr) Evar.Map.t
exception CannotFilter
-let filtering env cv_pb c1 c2 =
+let filtering sigma env cv_pb c1 c2 =
+ let open EConstr in
+ let open Vars in
let evm = ref Evar.Map.empty in
let define cv_pb e1 ev c1 =
try let (e2,c2) = Evar.Map.find ev !evm in
let shift = List.length e1 - List.length e2 in
- if constr_cmp cv_pb c1 (lift shift c2) then () else raise CannotFilter
+ if constr_cmp sigma cv_pb c1 (lift shift c2) then () else raise CannotFilter
with Not_found ->
evm := Evar.Map.add ev (e1,c1) !evm
in
let rec aux env cv_pb c1 c2 =
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind sigma c1, EConstr.kind sigma c2 with
| App _, App _ ->
- let ((p1,l1),(p2,l2)) = (split_app c1),(split_app c2) in
+ let ((p1,l1),(p2,l2)) = (split_app sigma c1),(split_app sigma c2) in
let () = aux env cv_pb l1 l2 in
begin match p1, p2 with
| [], [] -> ()
| (h1 :: p1), (h2 :: p2) ->
- aux env cv_pb (applistc h1 p1) (applistc h2 p2)
+ aux env cv_pb (applist (h1, p1)) (applist (h2, p2))
| _ -> assert false
end
| Prod (n,t1,c1), Prod (_,t2,c2) ->
@@ -875,52 +1251,52 @@ let filtering env cv_pb c1 c2 =
| _, Evar (ev,_) -> define cv_pb env ev c1
| Evar (ev,_), _ -> define cv_pb env ev c2
| _ ->
- if compare_constr_univ
+ if compare_constr_univ sigma
(fun pb c1 c2 -> aux env pb c1 c2; true) cv_pb c1 c2 then ()
else raise CannotFilter
(* TODO: le reste des binders *)
in
aux env cv_pb c1 c2; !evm
-let decompose_prod_letin : constr -> int * Context.Rel.t * constr =
- let rec prodec_rec i l c = match kind_of_term c with
+let decompose_prod_letin sigma c =
+ let rec prodec_rec i l c = match EConstr.kind sigma c with
| Prod (n,t,c) -> prodec_rec (succ i) (RelDecl.LocalAssum (n,t)::l) c
| LetIn (n,d,t,c) -> prodec_rec (succ i) (RelDecl.LocalDef (n,d,t)::l) c
| Cast (c,_,_) -> prodec_rec i l c
| _ -> i,l,c in
- prodec_rec 0 []
+ prodec_rec 0 [] c
(* (nb_lam [na1:T1]...[nan:Tan]c) where c is not an abstraction
* gives n (casts are ignored) *)
-let nb_lam =
- let rec nbrec n c = match kind_of_term c with
+let nb_lam sigma c =
+ let rec nbrec n c = match EConstr.kind sigma c with
| Lambda (_,_,c) -> nbrec (n+1) c
| Cast (c,_,_) -> nbrec n c
| _ -> n
in
- nbrec 0
+ nbrec 0 c
(* similar to nb_lam, but gives the number of products instead *)
-let nb_prod =
- let rec nbrec n c = match kind_of_term c with
+let nb_prod sigma c =
+ let rec nbrec n c = match EConstr.kind sigma c with
| Prod (_,_,c) -> nbrec (n+1) c
| Cast (c,_,_) -> nbrec n c
| _ -> n
in
- nbrec 0
+ nbrec 0 c
-let nb_prod_modulo_zeta x =
+let nb_prod_modulo_zeta sigma x =
let rec count n c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
Prod(_,_,t) -> count (n+1) t
- | LetIn(_,a,_,t) -> count n (subst1 a t)
+ | LetIn(_,a,_,t) -> count n (EConstr.Vars.subst1 a t)
| Cast(c,_,_) -> count n c
| _ -> n
in count 0 x
-let align_prod_letin c a : Context.Rel.t * constr =
- let (lc,_,_) = decompose_prod_letin c in
- let (la,l,a) = decompose_prod_letin a in
+let align_prod_letin sigma c a =
+ let (lc,_,_) = decompose_prod_letin sigma c in
+ let (la,l,a) = decompose_prod_letin sigma a in
if not (la >= lc) then invalid_arg "align_prod_letin";
let (l1,l2) = Util.List.chop lc l in
l2,it_mkProd_or_LetIn a l1
@@ -929,21 +1305,23 @@ let align_prod_letin c a : Context.Rel.t * constr =
(* [x1:c1;...;xn:cn]@(f;a1...an;x1;...;xn) --> @(f;a1...an) *)
(* Remplace 2 earlier buggish versions *)
-let rec eta_reduce_head c =
- match kind_of_term c with
+let rec eta_reduce_head sigma c =
+ let open EConstr in
+ let open Vars in
+ match EConstr.kind sigma c with
| Lambda (_,c1,c') ->
- (match kind_of_term (eta_reduce_head c') with
+ (match EConstr.kind sigma (eta_reduce_head sigma c') with
| App (f,cl) ->
let lastn = (Array.length cl) - 1 in
if lastn < 0 then anomaly (Pp.str "application without arguments")
else
- (match kind_of_term cl.(lastn) with
+ (match EConstr.kind sigma cl.(lastn) with
| Rel 1 ->
let c' =
if Int.equal lastn 0 then f
else mkApp (f, Array.sub cl 0 lastn)
in
- if noccurn 1 c'
+ if noccurn sigma 1 c'
then lift (-1) c'
else c
| _ -> c)
@@ -954,7 +1332,7 @@ let rec eta_reduce_head c =
(* iterator on rel context *)
let process_rel_context f env =
let sign = named_context_val env in
- let rels = rel_context env in
+ let rels = EConstr.rel_context env in
let env0 = reset_with_named_context sign env in
Context.Rel.fold_outside f rels ~init:env0
@@ -1002,6 +1380,14 @@ let fold_named_context_both_sides f l ~init = List.fold_right_and_left f l init
let mem_named_context_val id ctxt =
try ignore(Environ.lookup_named_val id ctxt); true with Not_found -> false
+let map_rel_decl f = function
+| RelDecl.LocalAssum (id, t) -> RelDecl.LocalAssum (id, f t)
+| RelDecl.LocalDef (id, b, t) -> RelDecl.LocalDef (id, f b, f t)
+
+let map_named_decl f = function
+| NamedDecl.LocalAssum (id, t) -> NamedDecl.LocalAssum (id, f t)
+| NamedDecl.LocalDef (id, b, t) -> NamedDecl.LocalDef (id, f b, f t)
+
let compact_named_context sign =
let compact l decl =
match decl, l with
@@ -1031,28 +1417,59 @@ let clear_named_body id env =
| d -> push_named d in
fold_named_context aux env ~init:(reset_context env)
-let global_vars env ids = Id.Set.elements (global_vars_set env ids)
+let global_vars_set env sigma constr =
+ let rec filtrec acc c =
+ let acc = match EConstr.kind sigma c with
+ | Var _ | Const _ | Ind _ | Construct _ ->
+ Id.Set.union (vars_of_global env (EConstr.to_constr sigma c)) acc
+ | _ -> acc
+ in
+ EConstr.fold sigma filtrec acc c
+ in
+ filtrec Id.Set.empty constr
+
+let global_vars env sigma ids = Id.Set.elements (global_vars_set env sigma ids)
-let global_vars_set_of_decl env = function
- | NamedDecl.LocalAssum (_,t) -> global_vars_set env t
+let global_vars_set_of_decl env sigma = function
+ | NamedDecl.LocalAssum (_,t) -> global_vars_set env sigma t
| NamedDecl.LocalDef (_,c,t) ->
- Id.Set.union (global_vars_set env t)
- (global_vars_set env c)
+ Id.Set.union (global_vars_set env sigma t)
+ (global_vars_set env sigma c)
-let dependency_closure env sign hyps =
+let dependency_closure env sigma sign hyps =
if Id.Set.is_empty hyps then [] else
let (_,lh) =
Context.Named.fold_inside
(fun (hs,hl) d ->
let x = NamedDecl.get_id d in
if Id.Set.mem x hs then
- (Id.Set.union (global_vars_set_of_decl env d) (Id.Set.remove x hs),
+ (Id.Set.union (global_vars_set_of_decl env sigma d) (Id.Set.remove x hs),
x::hl)
else (hs,hl))
~init:(hyps,[])
sign in
List.rev lh
+let global_app_of_constr sigma c =
+ let open Univ in
+ let open Globnames in
+ match EConstr.kind sigma c with
+ | Const (c, u) -> (ConstRef c, u), None
+ | Ind (i, u) -> (IndRef i, u), None
+ | Construct (c, u) -> (ConstructRef c, u), None
+ | Var id -> (VarRef id, EConstr.EInstance.empty), None
+ | Proj (p, c) -> (ConstRef (Projection.constant p), EConstr.EInstance.empty), Some c
+ | _ -> raise Not_found
+
+let prod_applist sigma c l =
+ let open EConstr in
+ let rec app subst c l =
+ match EConstr.kind sigma c, l with
+ | Prod(_,_,c), arg::l -> app (arg::subst) c l
+ | _, [] -> Vars.substl subst c
+ | _ -> anomaly (Pp.str "Not enough prod's") in
+ app [] c l
+
(* Combinators on judgments *)
let on_judgment f j = { uj_val = f j.uj_val; uj_type = f j.uj_type }
@@ -1071,6 +1488,7 @@ let context_chop k ctx =
(* Do not skip let-in's *)
let env_rel_context_chop k env =
+ let open EConstr in
let rels = rel_context env in
let ctx1,ctx2 = List.chop k rels in
push_rel_context ctx2 (reset_with_named_context (named_context_val env) env),
@@ -1084,6 +1502,7 @@ let impossible_default_case = ref None
let set_impossible_default_clause c = impossible_default_case := Some c
let coq_unit_judge =
+ let make_judge c t = make_judge (EConstr.of_constr c) (EConstr.of_constr t) in
let na1 = Name (Id.of_string "A") in
let na2 = Name (Id.of_string "H") in
fun () ->
diff --git a/engine/termops.mli b/engine/termops.mli
index 78826f79a..fe6dfb0ce 100644
--- a/engine/termops.mli
+++ b/engine/termops.mli
@@ -13,63 +13,54 @@ open Pp
open Names
open Term
open Environ
+open EConstr
(** printers *)
val print_sort : sorts -> std_ppcmds
val pr_sort_family : sorts_family -> std_ppcmds
-val pr_fix : (constr -> std_ppcmds) -> fixpoint -> std_ppcmds
-
-(** debug printer: do not use to display terms to the casual user... *)
-val set_print_constr : (env -> constr -> std_ppcmds) -> unit
-val print_constr : constr -> std_ppcmds
-val print_constr_env : env -> constr -> std_ppcmds
-val print_named_context : env -> std_ppcmds
-val pr_rel_decl : env -> Context.Rel.Declaration.t -> std_ppcmds
-val print_rel_context : env -> std_ppcmds
-val print_env : env -> std_ppcmds
+val pr_fix : ('a -> std_ppcmds) -> ('a, 'a) pfixpoint -> std_ppcmds
(** about contexts *)
val push_rel_assum : Name.t * types -> env -> env
-val push_rels_assum : (Name.t * types) list -> env -> env
-val push_named_rec_types : Name.t array * types array * 'a -> env -> env
+val push_rels_assum : (Name.t * Constr.types) list -> env -> env
+val push_named_rec_types : Name.t array * Constr.types array * 'a -> env -> env
-val lookup_rel_id : Id.t -> Context.Rel.t -> int * constr option * types
+val lookup_rel_id : Id.t -> ('c, 't) Context.Rel.pt -> int * 'c option * 't
(** Associates the contents of an identifier in a [rel_context]. Raise
[Not_found] if there is no such identifier. *)
(** Functions that build argument lists matching a block of binders or a context.
[rel_vect n m] builds [|Rel (n+m);...;Rel(n+1)|]
*)
-val rel_vect : int -> int -> constr array
+val rel_vect : int -> int -> Constr.constr array
val rel_list : int -> int -> constr list
(** iterators/destructors on terms *)
-val mkProd_or_LetIn : Context.Rel.Declaration.t -> types -> types
-val mkProd_wo_LetIn : Context.Rel.Declaration.t -> types -> types
+val mkProd_or_LetIn : rel_declaration -> types -> types
+val mkProd_wo_LetIn : rel_declaration -> types -> types
val it_mkProd : types -> (Name.t * types) list -> types
val it_mkLambda : constr -> (Name.t * types) list -> constr
-val it_mkProd_or_LetIn : types -> Context.Rel.t -> types
-val it_mkProd_wo_LetIn : types -> Context.Rel.t -> types
-val it_mkLambda_or_LetIn : constr -> Context.Rel.t -> constr
-val it_mkNamedProd_or_LetIn : types -> Context.Named.t -> types
-val it_mkNamedProd_wo_LetIn : types -> Context.Named.t -> types
-val it_mkNamedLambda_or_LetIn : constr -> Context.Named.t -> constr
+val it_mkProd_or_LetIn : types -> rel_context -> types
+val it_mkProd_wo_LetIn : types -> rel_context -> types
+val it_mkLambda_or_LetIn : Constr.constr -> Context.Rel.t -> Constr.constr
+val it_mkNamedProd_or_LetIn : types -> named_context -> types
+val it_mkNamedProd_wo_LetIn : Constr.types -> Context.Named.t -> Constr.types
+val it_mkNamedLambda_or_LetIn : constr -> named_context -> constr
(* Ad hoc version reinserting letin, assuming the body is defined in
the context where the letins are expanded *)
-val it_mkLambda_or_LetIn_from_no_LetIn : constr -> Context.Rel.t -> constr
+val it_mkLambda_or_LetIn_from_no_LetIn : Constr.constr -> Context.Rel.t -> Constr.constr
(** {6 Generic iterators on constr} *)
-val map_constr_with_named_binders :
- (Name.t -> 'a -> 'a) ->
- ('a -> constr -> constr) -> 'a -> constr -> constr
val map_constr_with_binders_left_to_right :
- (Context.Rel.Declaration.t -> 'a -> 'a) ->
+ Evd.evar_map ->
+ (rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) ->
'a -> constr -> constr
val map_constr_with_full_binders :
- (Context.Rel.Declaration.t -> 'a -> 'a) ->
+ Evd.evar_map ->
+ (rel_declaration -> 'a -> 'a) ->
('a -> constr -> constr) -> 'a -> constr -> constr
(** [fold_constr_with_binders g f n acc c] folds [f n] on the immediate
@@ -79,57 +70,58 @@ val map_constr_with_full_binders :
index) which is processed by [g] (which typically add 1 to [n]) at
each binder traversal; it is not recursive *)
-val fold_constr_with_binders :
+val fold_constr_with_binders : Evd.evar_map ->
('a -> 'a) -> ('a -> 'b -> constr -> 'b) -> 'a -> 'b -> constr -> 'b
-val fold_constr_with_full_binders :
+val fold_constr_with_full_binders : Evd.evar_map ->
(Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> 'b -> constr -> 'b) ->
'a -> 'b -> constr -> 'b
val iter_constr_with_full_binders :
- (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> constr -> unit) -> 'a ->
- constr -> unit
+ (Context.Rel.Declaration.t -> 'a -> 'a) -> ('a -> Constr.constr -> unit) -> 'a ->
+ Constr.constr -> unit
(**********************************************************************)
-val strip_head_cast : constr -> constr
-val drop_extra_implicit_args : constr -> constr
+val strip_head_cast : Evd.evar_map -> constr -> constr
+val drop_extra_implicit_args : Evd.evar_map -> constr -> constr
(** occur checks *)
exception Occur
-val occur_meta : types -> bool
-val occur_existential : types -> bool
-val occur_meta_or_existential : types -> bool
-val occur_evar : existential_key -> types -> bool
-val occur_var : env -> Id.t -> types -> bool
+val occur_meta : Evd.evar_map -> constr -> bool
+val occur_existential : Evd.evar_map -> constr -> bool
+val occur_meta_or_existential : Evd.evar_map -> constr -> bool
+val occur_evar : Evd.evar_map -> existential_key -> constr -> bool
+val occur_var : env -> Evd.evar_map -> Id.t -> constr -> bool
val occur_var_in_decl :
- env ->
- Id.t -> Context.Named.Declaration.t -> bool
+ env -> Evd.evar_map ->
+ Id.t -> named_declaration -> bool
(** As {!occur_var} but assume the identifier not to be a section variable *)
-val local_occur_var : Id.t -> types -> bool
+val local_occur_var : Evd.evar_map -> Id.t -> constr -> bool
-val free_rels : constr -> Int.Set.t
+val free_rels : Evd.evar_map -> constr -> Int.Set.t
(** [dependent m t] tests whether [m] is a subterm of [t] *)
-val dependent : constr -> constr -> bool
-val dependent_no_evar : constr -> constr -> bool
-val dependent_univs : constr -> constr -> bool
-val dependent_univs_no_evar : constr -> constr -> bool
-val dependent_in_decl : constr -> Context.Named.Declaration.t -> bool
-val count_occurrences : constr -> constr -> int
-val collect_metas : constr -> int list
-val collect_vars : constr -> Id.Set.t (** for visible vars only *)
+val dependent : Evd.evar_map -> constr -> constr -> bool
+val dependent_no_evar : Evd.evar_map -> constr -> constr -> bool
+val dependent_univs : Evd.evar_map -> constr -> constr -> bool
+val dependent_univs_no_evar : Evd.evar_map -> constr -> constr -> bool
+val dependent_in_decl : Evd.evar_map -> constr -> named_declaration -> bool
+val count_occurrences : Evd.evar_map -> constr -> constr -> int
+val collect_metas : Evd.evar_map -> constr -> int list
+val collect_vars : Evd.evar_map -> constr -> Id.Set.t (** for visible vars only *)
val vars_of_global_reference : env -> Globnames.global_reference -> Id.Set.t
-val occur_term : constr -> constr -> bool (** Synonymous
- of dependent
- Substitution of metavariables *)
-type meta_value_map = (metavariable * constr) list
-val subst_meta : meta_value_map -> constr -> constr
+val occur_term : Evd.evar_map -> constr -> constr -> bool (** Synonymous of dependent *)
+
+(* Substitution of metavariables *)
+type meta_value_map = (metavariable * Constr.constr) list
+val subst_meta : meta_value_map -> Constr.constr -> Constr.constr
+val isMetaOf : Evd.evar_map -> metavariable -> constr -> bool
(** Type assignment for metavariables *)
-type meta_type_map = (metavariable * types) list
+type meta_type_map = (metavariable * Constr.types) list
(** [pop c] lifts by -1 the positive indexes in [c] *)
val pop : constr -> constr
@@ -140,36 +132,38 @@ val pop : constr -> constr
(** [subst_term_gen eq d c] replaces [d] by [Rel 1] in [c] using [eq]
as equality *)
-val subst_term_gen :
- (constr -> constr -> bool) -> constr -> constr -> constr
+val subst_term_gen : Evd.evar_map ->
+ (Evd.evar_map -> constr -> constr -> bool) -> constr -> constr -> constr
(** [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq]
as equality *)
val replace_term_gen :
- (constr -> constr -> bool) ->
+ Evd.evar_map -> (Evd.evar_map -> constr -> constr -> bool) ->
constr -> constr -> constr -> constr
(** [subst_term d c] replaces [d] by [Rel 1] in [c] *)
-val subst_term : constr -> constr -> constr
+val subst_term : Evd.evar_map -> constr -> constr -> constr
(** [replace_term d e c] replaces [d] by [e] in [c] *)
-val replace_term : constr -> 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 compare_constr_univ : (Reduction.conv_pb -> constr -> constr -> bool) ->
+val compare_constr_univ : Evd.evar_map -> (Reduction.conv_pb -> constr -> constr -> bool) ->
Reduction.conv_pb -> constr -> constr -> bool
-val constr_cmp : Reduction.conv_pb -> constr -> constr -> bool
-val eq_constr : constr -> constr -> bool (* FIXME rename: erases universes*)
+val constr_cmp : Evd.evar_map -> Reduction.conv_pb -> constr -> constr -> bool
+val eq_constr : Evd.evar_map -> constr -> constr -> bool (* FIXME rename: erases universes*)
-val eta_reduce_head : constr -> constr
+val eta_reduce_head : Evd.evar_map -> constr -> constr
(** Flattens application lists *)
-val collapse_appl : constr -> constr
+val collapse_appl : Evd.evar_map -> constr -> constr
+
+val prod_applist : Evd.evar_map -> constr -> constr list -> constr
(** Remove recursively the casts around a term i.e.
[strip_outer_cast (Cast (Cast ... (Cast c, t) ... ))] is [c]. *)
-val strip_outer_cast : constr -> constr
+val strip_outer_cast : Evd.evar_map -> constr -> constr
exception CannotFilter
@@ -179,27 +173,27 @@ exception CannotFilter
(context,term), or raises [CannotFilter].
Warning: Outer-kernel sort subtyping are taken into account: c1 has
to be smaller than c2 wrt. sorts. *)
-type subst = (Context.Rel.t * constr) Evar.Map.t
-val filtering : Context.Rel.t -> Reduction.conv_pb -> constr -> constr -> subst
+type subst = (rel_context * constr) Evar.Map.t
+val filtering : Evd.evar_map -> rel_context -> Reduction.conv_pb -> constr -> constr -> subst
-val decompose_prod_letin : constr -> int * Context.Rel.t * constr
-val align_prod_letin : constr -> constr -> Context.Rel.t * constr
+val decompose_prod_letin : Evd.evar_map -> constr -> int * rel_context * constr
+val align_prod_letin : Evd.evar_map -> constr -> constr -> rel_context * constr
(** [nb_lam] {% $ %}[x_1:T_1]...[x_n:T_n]c{% $ %} where {% $ %}c{% $ %} is not an abstraction
gives {% $ %}n{% $ %} (casts are ignored) *)
-val nb_lam : constr -> int
+val nb_lam : Evd.evar_map -> constr -> int
(** Similar to [nb_lam], but gives the number of products instead *)
-val nb_prod : constr -> int
+val nb_prod : Evd.evar_map -> constr -> int
(** Similar to [nb_prod], but zeta-contracts let-in on the way *)
-val nb_prod_modulo_zeta : constr -> int
+val nb_prod_modulo_zeta : Evd.evar_map -> constr -> int
(** Get the last arg of a constr intended to be an application *)
-val last_arg : constr -> constr
+val last_arg : Evd.evar_map -> constr -> constr
(** Force the decomposition of a term as an applicative one *)
-val decompose_app_vect : constr -> constr * constr array
+val decompose_app_vect : Evd.evar_map -> constr -> constr * constr array
val adjust_app_list_size : constr -> constr list -> constr -> constr list ->
(constr * constr list * constr * constr list)
@@ -212,8 +206,8 @@ val add_name : Name.t -> names_context -> names_context
val lookup_name_of_rel : int -> names_context -> Name.t
val lookup_rel_of_name : Id.t -> names_context -> int
val empty_names_context : names_context
-val ids_of_rel_context : Context.Rel.t -> Id.t list
-val ids_of_named_context : Context.Named.t -> Id.t list
+val ids_of_rel_context : ('c, 't) Context.Rel.pt -> Id.t list
+val ids_of_named_context : ('c, 't) Context.Named.pt -> Id.t list
val ids_of_context : env -> Id.t list
val names_of_rel_context : env -> names_context
@@ -225,51 +219,89 @@ val context_chop : int -> Context.Rel.t -> Context.Rel.t * Context.Rel.t
(* [env_rel_context_chop n env] extracts out the [n] top declarations
of the rel_context part of [env], counting both local definitions and
hypotheses *)
-val env_rel_context_chop : int -> env -> env * Context.Rel.t
+val env_rel_context_chop : int -> env -> env * rel_context
(** Set of local names *)
val vars_of_env: env -> Id.Set.t
val add_vname : Id.Set.t -> Name.t -> Id.Set.t
(** other signature iterators *)
-val process_rel_context : (Context.Rel.Declaration.t -> env -> env) -> env -> env
-val assums_of_rel_context : Context.Rel.t -> (Name.t * constr) list
+val process_rel_context : (rel_declaration -> env -> env) -> env -> env
+val assums_of_rel_context : ('c, 't) Context.Rel.pt -> (Name.t * 't) list
val lift_rel_context : int -> Context.Rel.t -> Context.Rel.t
-val substl_rel_context : constr list -> Context.Rel.t -> Context.Rel.t
+val substl_rel_context : Constr.constr list -> Context.Rel.t -> Context.Rel.t
val smash_rel_context : Context.Rel.t -> Context.Rel.t (** expand lets in context *)
val map_rel_context_in_env :
- (env -> constr -> constr) -> env -> Context.Rel.t -> Context.Rel.t
+ (env -> Constr.constr -> Constr.constr) -> env -> Context.Rel.t -> Context.Rel.t
val map_rel_context_with_binders :
- (int -> constr -> constr) -> Context.Rel.t -> Context.Rel.t
+ (int -> 'c -> 'c) -> ('c, 'c) Context.Rel.pt -> ('c, 'c) Context.Rel.pt
val fold_named_context_both_sides :
('a -> Context.Named.Declaration.t -> Context.Named.Declaration.t list -> 'a) ->
Context.Named.t -> init:'a -> 'a
val mem_named_context_val : Id.t -> named_context_val -> bool
val compact_named_context : Context.Named.t -> Context.Compacted.t
+val map_rel_decl : ('a -> 'b) -> ('a, 'a) Context.Rel.Declaration.pt -> ('b, 'b) Context.Rel.Declaration.pt
+val map_named_decl : ('a -> 'b) -> ('a, 'a) Context.Named.Declaration.pt -> ('b, 'b) Context.Named.Declaration.pt
+
val clear_named_body : Id.t -> env -> env
-val global_vars : env -> constr -> Id.t list
-val global_vars_set_of_decl : env -> Context.Named.Declaration.t -> Id.Set.t
+val global_vars : env -> Evd.evar_map -> constr -> Id.t list
+val global_vars_set : env -> Evd.evar_map -> constr -> Id.Set.t
+val global_vars_set_of_decl : env -> Evd.evar_map -> named_declaration -> Id.Set.t
+val global_app_of_constr : Evd.evar_map -> constr -> (Globnames.global_reference * EInstance.t) * constr option
(** Gives an ordered list of hypotheses, closed by dependencies,
containing a given set *)
-val dependency_closure : env -> Context.Named.t -> Id.Set.t -> Id.t list
+val dependency_closure : env -> Evd.evar_map -> named_context -> Id.Set.t -> Id.t list
(** Test if an identifier is the basename of a global reference *)
val is_section_variable : Id.t -> bool
-val isGlobalRef : constr -> bool
+val global_of_constr : Evd.evar_map -> constr -> Globnames.global_reference * EInstance.t
+
+val is_global : Evd.evar_map -> Globnames.global_reference -> constr -> bool
+
+val isGlobalRef : Evd.evar_map -> constr -> bool
+
+val is_template_polymorphic : env -> Evd.evar_map -> constr -> bool
-val is_template_polymorphic : env -> constr -> bool
+val is_Prop : Evd.evar_map -> constr -> bool
(** Combinators on judgments *)
-val on_judgment : (types -> types) -> unsafe_judgment -> unsafe_judgment
-val on_judgment_value : (types -> types) -> unsafe_judgment -> unsafe_judgment
-val on_judgment_type : (types -> types) -> unsafe_judgment -> unsafe_judgment
+val on_judgment : ('a -> 'b) -> ('a, 'a) punsafe_judgment -> ('b, 'b) punsafe_judgment
+val on_judgment_value : ('c -> 'c) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment
+val on_judgment_type : ('t -> 't) -> ('c, 't) punsafe_judgment -> ('c, 't) punsafe_judgment
(** {6 Functions to deal with impossible cases } *)
-val set_impossible_default_clause : (unit -> (constr * types) Univ.in_universe_context_set) -> unit
+val set_impossible_default_clause : (unit -> (Constr.constr * Constr.types) Univ.in_universe_context_set) -> unit
val coq_unit_judge : unit -> unsafe_judgment Univ.in_universe_context_set
+
+(** {5 Debug pretty-printers} *)
+
+open Evd
+
+val pr_existential_key : evar_map -> evar -> Pp.std_ppcmds
+
+val pr_evar_suggested_name : existential_key -> evar_map -> Id.t
+
+val pr_evar_info : evar_info -> Pp.std_ppcmds
+val pr_evar_constraints : evar_map -> evar_constraint list -> Pp.std_ppcmds
+val pr_evar_map : ?with_univs:bool -> int option -> evar_map -> Pp.std_ppcmds
+val pr_evar_map_filter : ?with_univs:bool -> (Evar.t -> evar_info -> bool) ->
+ evar_map -> Pp.std_ppcmds
+val pr_metaset : Metaset.t -> Pp.std_ppcmds
+val pr_evar_universe_context : evar_universe_context -> Pp.std_ppcmds
+val pr_evd_level : evar_map -> Univ.Level.t -> Pp.std_ppcmds
+
+(** debug printer: do not use to display terms to the casual user... *)
+
+val set_print_constr : (env -> Evd.evar_map -> constr -> std_ppcmds) -> unit
+val print_constr : constr -> std_ppcmds
+val print_constr_env : env -> Evd.evar_map -> constr -> std_ppcmds
+val print_named_context : env -> std_ppcmds
+val pr_rel_decl : env -> Context.Rel.Declaration.t -> std_ppcmds
+val print_rel_context : env -> std_ppcmds
+val print_env : env -> std_ppcmds
diff --git a/engine/universes.ml b/engine/universes.ml
index 30a9ef163..ad5ff827b 100644
--- a/engine/universes.ml
+++ b/engine/universes.ml
@@ -445,15 +445,6 @@ let global_of_constr c =
| Var id -> VarRef id, Instance.empty
| _ -> raise Not_found
-let global_app_of_constr c =
- match kind_of_term c with
- | Const (c, u) -> (ConstRef c, u), None
- | Ind (i, u) -> (IndRef i, u), None
- | Construct (c, u) -> (ConstructRef c, u), None
- | Var id -> (VarRef id, Instance.empty), None
- | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.empty), Some c
- | _ -> raise Not_found
-
open Declarations
let type_of_reference env r =
diff --git a/engine/universes.mli b/engine/universes.mli
index 725c21d29..932de941a 100644
--- a/engine/universes.mli
+++ b/engine/universes.mli
@@ -76,8 +76,8 @@ val eq_constr_univs_infer : UGraph.t -> 'a constraint_accumulator ->
{!eq_constr_univs_infer} taking kind-of-term functions, to expose
subterms of [m] and [n], arguments. *)
val eq_constr_univs_infer_with :
- (constr -> (constr,types) kind_of_term) ->
- (constr -> (constr,types) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
UGraph.t -> 'a constraint_accumulator -> constr -> constr -> 'a -> 'a option
(** [leq_constr_univs u a b] is [true, c] if [a] is convertible to [b]
@@ -130,8 +130,6 @@ val fresh_universe_context_set_instance : universe_context_set ->
(** Raises [Not_found] if not a global reference. *)
val global_of_constr : constr -> Globnames.global_reference puniverses
-val global_app_of_constr : constr -> Globnames.global_reference puniverses * constr option
-
val constr_of_global_univ : Globnames.global_reference puniverses -> constr
val extend_context : 'a in_universe_context_set -> universe_context_set ->
diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml
index 8cadf1a26..0dfa03cca 100644
--- a/ide/ide_slave.ml
+++ b/ide/ide_slave.ml
@@ -193,7 +193,7 @@ let process_goal sigma g =
pr_goal_concl_style_env env sigma norm_constr
in
let process_hyp d (env,l) =
- let d = CompactedDecl.map_constr (Reductionops.nf_evar sigma) d in
+ let d = CompactedDecl.map_constr (fun c -> EConstr.Unsafe.to_constr (Reductionops.nf_evar sigma (EConstr.of_constr c))) d in
let d' = CompactedDecl.to_named_context d in
(List.fold_right Environ.push_named d' env,
(pr_compacted_decl env sigma d) :: l) in
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 925e9517c..f272d219a 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -961,6 +961,7 @@ let extern_constr_gen lax goal_concl_style scopt env sigma t =
(* Not "goal_concl_style" means do alpha-conversion avoiding only *)
(* those goal/section/rel variables that occurs in the subterm under *)
(* consideration; see namegen.ml for further details *)
+ let t = EConstr.of_constr t in
let avoid = if goal_concl_style then ids_of_context env else [] in
let r = Detyping.detype ~lax:lax goal_concl_style avoid env sigma t in
let vars = vars_of_env env in
@@ -973,6 +974,7 @@ let extern_constr ?(lax=false) goal_concl_style env sigma t =
extern_constr_gen lax goal_concl_style None env sigma t
let extern_type goal_concl_style env sigma t =
+ let t = EConstr.of_constr t in
let avoid = if goal_concl_style then ids_of_context env else [] in
let r = Detyping.detype goal_concl_style avoid env sigma t in
extern_glob_type (vars_of_env env) r
@@ -1051,14 +1053,16 @@ let rec glob_of_pat env sigma = function
| _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive")
in
GCases (loc,RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat)
- | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (mkFix f) (** FIXME bad env *)
- | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (mkCoFix c)
+ | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *)
+ | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c))
| PSort s -> GSort (loc,s)
let extern_constr_pattern env sigma pat =
extern true (None,[]) Id.Set.empty (glob_of_pat env sigma pat)
let extern_rel_context where env sigma sign =
+ let sign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) sign in
+ let where = Option.map EConstr.of_constr where in
let a = detype_rel_context where [] (names_of_rel_context env,env) sigma sign in
let vars = vars_of_env env in
let a = List.map (extended_glob_local_binder_of_decl Loc.ghost) a in
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 8fe6ce85e..d75487ecf 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1831,7 +1831,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(add_name match_acc (loc,x)) ((loc,x)::var_acc)
| (LocalAssum (cano_name,ty) | LocalDef (cano_name,_,ty)) :: t, c::tt ->
let fresh =
- Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names ty in
+ Namegen.next_name_away_with_default_using_types "iV" cano_name forbidden_names (EConstr.of_constr ty) in
canonize_args t tt (fresh::forbidden_names)
((fresh,c)::match_acc) ((cases_pattern_loc c,Name fresh)::var_acc)
| _ -> assert false in
@@ -1925,7 +1925,7 @@ let extract_ids env =
let scope_of_type_kind = function
| IsType -> Notation.current_type_scope_name ()
- | OfType typ -> compute_type_scope typ
+ | OfType typ -> compute_type_scope (EConstr.Unsafe.to_constr typ)
| WithoutTypeConstraint -> None
let empty_ltac_sign = {
@@ -2004,7 +2004,7 @@ let interp_constr_evars env evdref ?(impls=empty_internalization_env) c =
interp_constr_evars_gen env evdref WithoutTypeConstraint ~impls c
let interp_casted_constr_evars env evdref ?(impls=empty_internalization_env) c typ =
- interp_constr_evars_gen env evdref ~impls (OfType typ) c
+ interp_constr_evars_gen env evdref ~impls (OfType (EConstr.of_constr typ)) c
let interp_type_evars env evdref ?(impls=empty_internalization_env) c =
interp_constr_evars_gen env evdref IsType ~impls c
@@ -2064,6 +2064,7 @@ let intern_context global_level env impl_env binders =
user_err ~loc ~hdr:"internalize" (explain_internalization_error e)
let interp_rawcontext_evars env evdref k bl =
+ let open EConstr in
let (env, par, _, impls) =
List.fold_left
(fun (env,params,n,impls) (na, k, b, t) ->
diff --git a/interp/constrintern.mli b/interp/constrintern.mli
index e45de2588..758d4e650 100644
--- a/interp/constrintern.mli
+++ b/interp/constrintern.mli
@@ -99,7 +99,7 @@ val interp_constr : env -> evar_map -> ?impls:internalization_env ->
constr_expr -> constr Evd.in_evar_universe_context
val interp_casted_constr : env -> evar_map -> ?impls:internalization_env ->
- constr_expr -> types -> constr Evd.in_evar_universe_context
+ constr_expr -> EConstr.types -> constr Evd.in_evar_universe_context
val interp_type : env -> evar_map -> ?impls:internalization_env ->
constr_expr -> types Evd.in_evar_universe_context
@@ -107,32 +107,32 @@ val interp_type : env -> evar_map -> ?impls:internalization_env ->
(** Main interpretation function expecting all postponed problems to
be resolved, but possibly leaving evars. *)
-val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * constr
+val interp_open_constr : env -> evar_map -> constr_expr -> evar_map * EConstr.constr
(** Accepting unresolved evars *)
val interp_constr_evars : env -> evar_map ref ->
- ?impls:internalization_env -> constr_expr -> constr
+ ?impls:internalization_env -> constr_expr -> EConstr.constr
val interp_casted_constr_evars : env -> evar_map ref ->
- ?impls:internalization_env -> constr_expr -> types -> constr
+ ?impls:internalization_env -> constr_expr -> types -> EConstr.constr
val interp_type_evars : env -> evar_map ref ->
- ?impls:internalization_env -> constr_expr -> types
+ ?impls:internalization_env -> constr_expr -> EConstr.types
(** Accepting unresolved evars and giving back the manual implicit arguments *)
val interp_constr_evars_impls : env -> evar_map ref ->
?impls:internalization_env -> constr_expr ->
- constr * Impargs.manual_implicits
+ EConstr.constr * Impargs.manual_implicits
val interp_casted_constr_evars_impls : env -> evar_map ref ->
- ?impls:internalization_env -> constr_expr -> types ->
- constr * Impargs.manual_implicits
+ ?impls:internalization_env -> constr_expr -> EConstr.types ->
+ EConstr.constr * Impargs.manual_implicits
val interp_type_evars_impls : env -> evar_map ref ->
?impls:internalization_env -> constr_expr ->
- types * Impargs.manual_implicits
+ EConstr.types * Impargs.manual_implicits
(** Interprets constr patterns *)
@@ -151,14 +151,14 @@ val interp_reference : ltac_sign -> reference -> glob_constr
val interp_binder : env -> evar_map -> Name.t -> constr_expr ->
types Evd.in_evar_universe_context
-val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> types
+val interp_binder_evars : env -> evar_map ref -> Name.t -> constr_expr -> EConstr.types
(** Interpret contexts: returns extended env and context *)
val interp_context_evars :
?global_level:bool -> ?impl_env:internalization_env -> ?shift:int ->
env -> evar_map ref -> local_binder_expr list ->
- internalization_env * ((env * Context.Rel.t) * Impargs.manual_implicits)
+ internalization_env * ((env * EConstr.rel_context) * Impargs.manual_implicits)
(* val interp_context_gen : (env -> glob_constr -> unsafe_type_judgment Evd.in_evar_universe_context) -> *)
(* (env -> Evarutil.type_constraint -> glob_constr -> unsafe_judgment Evd.in_evar_universe_context) -> *)
@@ -175,7 +175,7 @@ val interp_context_evars :
val locate_reference : Libnames.qualid -> Globnames.global_reference
val is_global : Id.t -> bool
-val construct_reference : Context.Named.t -> Id.t -> constr
+val construct_reference : ('c, 't) Context.Named.pt -> Id.t -> constr
val global_reference : Id.t -> constr
val global_reference_in_absolute_module : DirPath.t -> Id.t -> constr
diff --git a/interp/notation.ml b/interp/notation.ml
index 66d3c9185..90ac7f729 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -585,7 +585,7 @@ let scope_class_compare : scope_class -> scope_class -> int =
cl_typ_ord
let compute_scope_class t =
- let (cl,_,_) = find_class_type Evd.empty t in
+ let (cl,_,_) = find_class_type Evd.empty (EConstr.of_constr t) in
cl
module ScopeClassOrd =
@@ -615,7 +615,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 (Reductionops.whd_betaiotazeta Evd.empty t) with
+ match kind_of_term (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_ops.ml b/interp/notation_ops.ml
index 59625426f..8b4fadb5a 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -442,6 +442,7 @@ let notation_constr_of_glob_constr nenv a =
(* Substitution of kernel names, avoiding a list of bound identifiers *)
let notation_constr_of_constr avoiding t =
+ let t = EConstr.of_constr t in
let t = Detyping.detype false avoiding (Global.env()) Evd.empty t in
let nenv = {
ninterp_var_type = Id.Map.empty;
diff --git a/interp/reserve.ml b/interp/reserve.ml
index a4d4f4027..1565ba4a9 100644
--- a/interp/reserve.ml
+++ b/interp/reserve.ml
@@ -107,7 +107,9 @@ let constr_key c =
let revert_reserved_type t =
try
+ let t = EConstr.Unsafe.to_constr t in
let reserved = KeyMap.find (constr_key t) !reserve_revtable in
+ let t = EConstr.of_constr t in
let t = Detyping.detype false [] (Global.env()) Evd.empty t in
(* pedrot: if [Notation_ops.match_notation_constr] may raise [Failure _]
then I've introduced a bug... *)
diff --git a/interp/stdarg.mli b/interp/stdarg.mli
index af3a73462..113fe40ba 100644
--- a/interp/stdarg.mli
+++ b/interp/stdarg.mli
@@ -11,6 +11,7 @@
open Loc
open Names
open Term
+open EConstr
open Libnames
open Globnames
open Genredexpr
diff --git a/intf/pattern.mli b/intf/pattern.mli
index a32e7e4b9..48381cacd 100644
--- a/intf/pattern.mli
+++ b/intf/pattern.mli
@@ -43,11 +43,11 @@ open Misctypes
could be inferred. We also loose the ability of typing ltac
variables before calling the right-hand-side of ltac matching clauses. *)
-type constr_under_binders = Id.t list * constr
+type constr_under_binders = Id.t list * EConstr.constr
(** Types of substitutions with or w/o bound variables *)
-type patvar_map = constr Id.Map.t
+type patvar_map = EConstr.constr Id.Map.t
type extended_patvar_map = constr_under_binders Id.Map.t
(** {5 Patterns} *)
diff --git a/intf/tactypes.mli b/intf/tactypes.mli
index b96cb67df..02cfc44e2 100644
--- a/intf/tactypes.mli
+++ b/intf/tactypes.mli
@@ -26,8 +26,8 @@ type glob_constr_pattern_and_expr = Id.Set.t * glob_constr_and_expr * constr_pat
type 'a delayed_open =
{ delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma }
-type delayed_open_constr = Term.constr delayed_open
-type delayed_open_constr_with_bindings = Term.constr with_bindings delayed_open
+type delayed_open_constr = EConstr.constr delayed_open
+type delayed_open_constr_with_bindings = EConstr.constr with_bindings delayed_open
type intro_pattern = delayed_open_constr intro_pattern_expr located
type intro_patterns = delayed_open_constr intro_pattern_expr located list
diff --git a/kernel/constr.ml b/kernel/constr.ml
index ce20751ab..5a7561bf5 100644
--- a/kernel/constr.ml
+++ b/kernel/constr.ml
@@ -81,27 +81,27 @@ type pconstructor = constructor puniverses
(* [Var] is used for named variables and [Rel] for variables as
de Bruijn indices. *)
-type ('constr, 'types) kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of Sorts.t
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of pconstant
- | Ind of pinductive
- | Construct of pconstructor
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
| Proj of projection * 'constr
(* constr is the fixpoint of the previous type. Requires option
-rectypes of the Caml compiler to be set *)
-type t = (t,t) kind_of_term
+type t = (t, t, Sorts.t, Instance.t) kind_of_term
type constr = t
type existential = existential_key * constr array
@@ -235,6 +235,12 @@ let mkVar id = Var id
let kind c = c
+(* The other way around. We treat specifically smart constructors *)
+let of_kind = function
+| App (f, a) -> mkApp (f, a)
+| Cast (c, knd, t) -> mkCast (c, knd, t)
+| k -> k
+
(****************************************************************************)
(* Functions to recur through subterms *)
(****************************************************************************)
diff --git a/kernel/constr.mli b/kernel/constr.mli
index 7095dbe6f..700c235e6 100644
--- a/kernel/constr.mli
+++ b/kernel/constr.mli
@@ -188,7 +188,7 @@ type ('constr, 'types) pfixpoint =
type ('constr, 'types) pcofixpoint =
int * ('constr, 'types) prec_declaration
-type ('constr, 'types) kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
| Rel of int (** Gallina-variable introduced by [forall], [fun], [let-in], [fix], or [cofix]. *)
| Var of Id.t (** Gallina-variable that was introduced by Vernacular-command that extends
@@ -197,7 +197,7 @@ type ('constr, 'types) kind_of_term =
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of Sorts.t
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types (** Concrete syntax ["forall A:B,C"] is represented as [Prod (A,B,C)]. *)
| Lambda of Name.t * 'types * 'constr (** Concrete syntax ["fun A:B => C"] is represented as [Lambda (A,B,C)]. *)
@@ -208,11 +208,11 @@ type ('constr, 'types) kind_of_term =
- [F] itself is not {!App}
- and [[|P1;..;Pn|]] is not empty. *)
- | Const of constant puniverses (** Gallina-variable that was introduced by Vernacular-command that extends the global environment
+ | Const of (constant * 'univs) (** Gallina-variable that was introduced by Vernacular-command that extends the global environment
(i.e. [Parameter], or [Axiom], or [Definition], or [Theorem] etc.) *)
- | Ind of inductive puniverses (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
- | Construct of constructor puniverses (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
+ | Ind of (inductive * 'univs) (** A name of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
+ | Construct of (constructor * 'univs) (** A constructor of an inductive type defined by [Variant], [Inductive] or [Record] Vernacular-commands. *)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
@@ -222,7 +222,8 @@ type ('constr, 'types) kind_of_term =
least one argument and the function is not itself an applicative
term *)
-val kind : constr -> (constr, types) kind_of_term
+val kind : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
+val of_kind : (constr, types, Sorts.t, Univ.Instance.t) kind_of_term -> constr
(** [equal a b] is true if [a] equals [b] modulo alpha, casts,
and application grouping *)
@@ -310,13 +311,22 @@ val compare_head_gen : (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
(constr -> constr -> bool) ->
constr -> constr -> bool
+val compare_head_gen_leq_with :
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
+ (Sorts.t -> Sorts.t -> bool) ->
+ (constr -> constr -> bool) ->
+ (constr -> constr -> bool) ->
+ constr -> constr -> bool
+
(** [compare_head_gen_with k1 k2 u s f c1 c2] compares [c1] and [c2]
like [compare_head_gen u s f c1 c2], except that [k1] (resp. [k2])
is used,rather than {!kind}, to expose the immediate subterms of
[c1] (resp. [c2]). *)
val compare_head_gen_with :
- (constr -> (constr,types) kind_of_term) ->
- (constr -> (constr,types) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
+ (constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term) ->
(bool -> Univ.Instance.t -> Univ.Instance.t -> bool) ->
(Sorts.t -> Sorts.t -> bool) ->
(constr -> constr -> bool) ->
diff --git a/kernel/context.ml b/kernel/context.ml
index ae0388003..abb284f22 100644
--- a/kernel/context.ml
+++ b/kernel/context.ml
@@ -37,9 +37,11 @@ struct
module Declaration =
struct
(* local declaration *)
- type t =
- | LocalAssum of Name.t * Constr.t (** name, type *)
- | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *)
+ type ('constr, 'types) pt =
+ | LocalAssum of Name.t * 'types (** name, type *)
+ | LocalDef of Name.t * 'constr * 'types (** name, value, type *)
+
+ type t = (Constr.constr, Constr.types) pt
(** Return the name bound by a given declaration. *)
let get_name = function
@@ -87,12 +89,12 @@ struct
| LocalDef (_, v, ty) -> f v && f ty
(** Check whether the two given declarations are equal. *)
- let equal decl1 decl2 =
+ let equal eq decl1 decl2 =
match decl1, decl2 with
| LocalAssum (n1,ty1), LocalAssum (n2, ty2) ->
- Name.equal n1 n2 && Constr.equal ty1 ty2
+ Name.equal n1 n2 && eq ty1 ty2
| LocalDef (n1,v1,ty1), LocalDef (n2,v2,ty2) ->
- Name.equal n1 n2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ Name.equal n1 n2 && eq v1 v2 && eq ty1 ty2
| _ ->
false
@@ -152,6 +154,7 @@ struct
(** Rel-context is represented as a list of declarations.
Inner-most declarations are at the beginning of the list.
Outer-most declarations are at the end of the list. *)
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
(** empty rel-context *)
@@ -166,14 +169,14 @@ struct
(** [extended_rel_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
with n = |Δ| and with the local definitions of [Γ] skipped in
[args]. Example: for [x:T,y:=c,z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
- let nhyps =
+ let nhyps ctx =
let open Declaration in
let rec nhyps acc = function
| [] -> acc
| LocalAssum _ :: hyps -> nhyps (succ acc) hyps
| LocalDef _ :: hyps -> nhyps acc hyps
in
- nhyps 0
+ nhyps 0 ctx
(** Return a declaration designated by a given de Bruijn index.
@raise Not_found if the designated de Bruijn index is not present in the designated rel-context. *)
@@ -184,7 +187,7 @@ struct
| _, [] -> raise Not_found
(** Check whether given two rel-contexts are equal. *)
- let equal = List.equal Declaration.equal
+ let equal eq l = List.equal (fun c -> Declaration.equal eq c) l
(** Map all terms in a given rel-context. *)
let map f = List.smartmap (Declaration.map_constr f)
@@ -202,26 +205,26 @@ struct
(** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
and each {e local definition} is mapped to [false]. *)
- let to_tags =
+ let to_tags l =
let rec aux l = function
| [] -> l
| Declaration.LocalDef _ :: ctx -> aux (true::l) ctx
| Declaration.LocalAssum _ :: ctx -> aux (false::l) ctx
- in aux []
+ in aux [] l
(** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
with n = |Δ| and with the {e local definitions} of [Γ] skipped in
[args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
- let to_extended_list n =
+ let to_extended_list mk n l =
let rec reln l p = function
- | Declaration.LocalAssum _ :: hyps -> reln (Constr.mkRel (n+p) :: l) (p+1) hyps
+ | Declaration.LocalAssum _ :: hyps -> reln (mk (n+p) :: l) (p+1) hyps
| Declaration.LocalDef _ :: hyps -> reln l (p+1) hyps
| [] -> l
in
- reln [] 1
+ reln [] 1 l
(** [extended_vect n Γ] does the same, returning instead an array. *)
- let to_extended_vect n hyps = Array.of_list (to_extended_list n hyps)
+ let to_extended_vect mk n hyps = Array.of_list (to_extended_list mk n hyps)
end
(** This module represents contexts that can capture non-anonymous variables.
@@ -232,9 +235,11 @@ struct
module Declaration =
struct
(** local declaration *)
- type t =
- | LocalAssum of Id.t * Constr.t (** identifier, type *)
- | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *)
+ type ('constr, 'types) pt =
+ | LocalAssum of Id.t * 'types (** identifier, type *)
+ | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *)
+
+ type t = (Constr.constr, Constr.types) pt
(** Return the identifier bound by a given declaration. *)
let get_id = function
@@ -282,12 +287,12 @@ struct
| LocalDef (_, v, ty) -> f v && f ty
(** Check whether the two given declarations are equal. *)
- let equal decl1 decl2 =
+ let equal eq decl1 decl2 =
match decl1, decl2 with
| LocalAssum (id1, ty1), LocalAssum (id2, ty2) ->
- Id.equal id1 id2 && Constr.equal ty1 ty2
+ Id.equal id1 id2 && eq ty1 ty2
| LocalDef (id1, v1, ty1), LocalDef (id2, v2, ty2) ->
- Id.equal id1 id2 && Constr.equal v1 v2 && Constr.equal ty1 ty2
+ Id.equal id1 id2 && eq v1 v2 && eq ty1 ty2
| _ ->
false
@@ -362,6 +367,7 @@ struct
(** Named-context is represented as a list of declarations.
Inner-most declarations are at the beginning of the list.
Outer-most declarations are at the end of the list. *)
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
(** empty named-context *)
@@ -380,7 +386,7 @@ struct
| [] -> raise Not_found
(** Check whether given two named-contexts are equal. *)
- let equal = List.equal Declaration.equal
+ let equal eq l = List.equal (fun c -> Declaration.equal eq c) l
(** Map all terms in a given named-context. *)
let map f = List.smartmap (Declaration.map_constr f)
@@ -397,28 +403,30 @@ struct
let fold_outside f l ~init = List.fold_right f l init
(** Return the set of all identifiers bound in a given named-context. *)
- let to_vars =
- List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty
+ let to_vars l =
+ List.fold_left (fun accu decl -> Id.Set.add (Declaration.get_id decl) accu) Id.Set.empty l
(** [instance_from_named_context Ω] builds an instance [args] such
that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local
definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it
gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
- let to_instance =
+ let to_instance mk l =
let filter = function
- | Declaration.LocalAssum (id, _) -> Some (Constr.mkVar id)
+ | Declaration.LocalAssum (id, _) -> Some (mk id)
| _ -> None
in
- List.map_filter filter
+ List.map_filter filter l
end
module Compacted =
struct
module Declaration =
struct
- type t =
- | LocalAssum of Id.t list * Constr.t
- | LocalDef of Id.t list * Constr.t * Constr.t
+ type ('constr, 'types) pt =
+ | LocalAssum of Id.t list * 'types
+ | LocalDef of Id.t list * 'constr * 'types
+
+ type t = (Constr.constr, Constr.types) pt
let map_constr f = function
| LocalAssum (ids, ty) as decl ->
@@ -442,6 +450,7 @@ module Compacted =
List.map (fun id -> Named.Declaration.LocalDef (id,v,t)) ids
end
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
let fold f l ~init = List.fold_right f l init
diff --git a/kernel/context.mli b/kernel/context.mli
index 955e214cb..0c666a25d 100644
--- a/kernel/context.mli
+++ b/kernel/context.mli
@@ -29,110 +29,115 @@ sig
module Declaration :
sig
(* local declaration *)
- type t = LocalAssum of Name.t * Constr.t (** name, type *)
- | LocalDef of Name.t * Constr.t * Constr.t (** name, value, type *)
+ type ('constr, 'types) pt =
+ | LocalAssum of Name.t * 'types (** name, type *)
+ | LocalDef of Name.t * 'constr * 'types (** name, value, type *)
+
+ type t = (Constr.constr, Constr.types) pt
(** Return the name bound by a given declaration. *)
- val get_name : t -> Name.t
+ val get_name : ('c, 't) pt -> Name.t
(** Return [Some value] for local-declarations and [None] for local-assumptions. *)
- val get_value : t -> Constr.t option
+ val get_value : ('c, 't) pt -> 'c option
(** Return the type of the name bound by a given declaration. *)
- val get_type : t -> Constr.t
+ val get_type : ('c, 't) pt -> 't
(** Set the name that is bound by a given declaration. *)
- val set_name : Name.t -> t -> t
+ val set_name : Name.t -> ('c, 't) pt -> ('c, 't) pt
(** Set the type of the bound variable in a given declaration. *)
- val set_type : Constr.t -> t -> t
+ val set_type : 't -> ('c, 't) pt -> ('c, 't) pt
(** Return [true] iff a given declaration is a local assumption. *)
- val is_local_assum : t -> bool
+ val is_local_assum : ('c, 't) pt -> bool
(** Return [true] iff a given declaration is a local definition. *)
- val is_local_def : t -> bool
+ val is_local_def : ('c, 't) pt -> bool
(** Check whether any term in a given declaration satisfies a given predicate. *)
- val exists : (Constr.t -> bool) -> t -> bool
+ val exists : ('c -> bool) -> ('c, 'c) pt -> bool
(** Check whether all terms in a given declaration satisfy a given predicate. *)
- val for_all : (Constr.t -> bool) -> t -> bool
+ val for_all : ('c -> bool) -> ('c, 'c) pt -> bool
(** Check whether the two given declarations are equal. *)
- val equal : t -> t -> bool
+ val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool
(** Map the name bound by a given declaration. *)
- val map_name : (Name.t -> Name.t) -> t -> t
+ val map_name : (Name.t -> Name.t) -> ('c, 't) pt -> ('c, 't) pt
(** For local assumptions, this function returns the original local assumptions.
For local definitions, this function maps the value in the local definition. *)
- val map_value : (Constr.t -> Constr.t) -> t -> t
+ val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt
(** Map the type of the name bound by a given declaration. *)
- val map_type : (Constr.t -> Constr.t) -> t -> t
+ val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt
(** Map all terms in a given declaration. *)
- val map_constr : (Constr.t -> Constr.t) -> t -> t
+ val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
(** Perform a given action on all terms in a given declaration. *)
- val iter_constr : (Constr.t -> unit) -> t -> unit
+ val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit
(** Reduce all terms in a given declaration to a single value. *)
- val fold_constr : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+ val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a
- val to_tuple : t -> Name.t * Constr.t option * Constr.t
+ val to_tuple : ('c, 't) pt -> Name.t * 'c option * 't
end
(** Rel-context is represented as a list of declarations.
Inner-most declarations are at the beginning of the list.
Outer-most declarations are at the end of the list. *)
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
(** empty rel-context *)
- val empty : t
+ val empty : ('c, 't) pt
(** Return a new rel-context enriched by with a given inner-most declaration. *)
- val add : Declaration.t -> t -> t
+ val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt
(** Return the number of {e local declarations} in a given context. *)
- val length : t -> int
+ val length : ('c, 't) pt -> int
(** Check whether given two rel-contexts are equal. *)
- val equal : t -> t -> bool
+ val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool
(** Return the number of {e local assumptions} in a given rel-context. *)
- val nhyps : t -> int
+ val nhyps : ('c, 't) pt -> int
(** Return a declaration designated by a given de Bruijn index.
@raise Not_found if the designated de Bruijn index outside the range. *)
- val lookup : int -> t -> Declaration.t
+ val lookup : int -> ('c, 't) pt -> ('c, 't) Declaration.pt
(** Map all terms in a given rel-context. *)
- val map : (Constr.t -> Constr.t) -> t -> t
+ val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
(** Perform a given action on every declaration in a given rel-context. *)
- val iter : (Constr.t -> unit) -> t -> unit
+ val iter : ('c -> unit) -> ('c, 'c) pt -> unit
(** Reduce all terms in a given rel-context to a single value.
Innermost declarations are processed first. *)
- val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
+ val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a
(** Reduce all terms in a given rel-context to a single value.
Outermost declarations are processed first. *)
- val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
+ val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a
(** Map a given rel-context to a list where each {e local assumption} is mapped to [true]
and each {e local definition} is mapped to [false]. *)
- val to_tags : t -> bool list
+ val to_tags : ('c, 't) pt -> bool list
- (** [extended_list n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
+ (** [extended_list mk n Γ] builds an instance [args] such that [Γ,Δ ⊢ args:Γ]
with n = |Δ| and with the {e local definitions} of [Γ] skipped in
- [args]. Example: for [x:T, y:=c, z:U] and [n]=2, it gives [Rel 5, Rel 3]. *)
- val to_extended_list : int -> t -> Constr.t list
+ [args] where [mk] is used to build the corresponding variables.
+ Example: for [x:T, y:=c, z:U] and [n]=2, it gives [mk 5, mk 3]. *)
+ val to_extended_list : (int -> 'r) -> int -> ('c, 't) pt -> 'r list
(** [extended_vect n Γ] does the same, returning instead an array. *)
- val to_extended_vect : int -> t -> Constr.t array
+ val to_extended_vect : (int -> 'r) -> int -> ('c, 't) pt -> 'r array
end
(** This module represents contexts that can capture non-anonymous variables.
@@ -142,129 +147,136 @@ sig
(** Representation of {e local declarations}. *)
module Declaration :
sig
- type t = LocalAssum of Id.t * Constr.t (** identifier, type *)
- | LocalDef of Id.t * Constr.t * Constr.t (** identifier, value, type *)
+ type ('constr, 'types) pt =
+ | LocalAssum of Id.t * 'types (** identifier, type *)
+ | LocalDef of Id.t * 'constr * 'types (** identifier, value, type *)
+
+ type t = (Constr.constr, Constr.types) pt
(** Return the identifier bound by a given declaration. *)
- val get_id : t -> Id.t
+ val get_id : ('c, 't) pt -> Id.t
(** Return [Some value] for local-declarations and [None] for local-assumptions. *)
- val get_value : t -> Constr.t option
+ val get_value : ('c, 't) pt -> 'c option
(** Return the type of the name bound by a given declaration. *)
- val get_type : t -> Constr.t
+ val get_type : ('c, 't) pt -> 't
(** Set the identifier that is bound by a given declaration. *)
- val set_id : Id.t -> t -> t
+ val set_id : Id.t -> ('c, 't) pt -> ('c, 't) pt
(** Set the type of the bound variable in a given declaration. *)
- val set_type : Constr.t -> t -> t
+ val set_type : 't -> ('c, 't) pt -> ('c, 't) pt
(** Return [true] iff a given declaration is a local assumption. *)
- val is_local_assum : t -> bool
+ val is_local_assum : ('c, 't) pt -> bool
(** Return [true] iff a given declaration is a local definition. *)
- val is_local_def : t -> bool
+ val is_local_def : ('c, 't) pt -> bool
(** Check whether any term in a given declaration satisfies a given predicate. *)
- val exists : (Constr.t -> bool) -> t -> bool
+ val exists : ('c -> bool) -> ('c, 'c) pt -> bool
(** Check whether all terms in a given declaration satisfy a given predicate. *)
- val for_all : (Constr.t -> bool) -> t -> bool
+ val for_all : ('c -> bool) -> ('c, 'c) pt -> bool
(** Check whether the two given declarations are equal. *)
- val equal : t -> t -> bool
+ val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool
(** Map the identifier bound by a given declaration. *)
- val map_id : (Id.t -> Id.t) -> t -> t
+ val map_id : (Id.t -> Id.t) -> ('c, 't) pt -> ('c, 't) pt
(** For local assumptions, this function returns the original local assumptions.
For local definitions, this function maps the value in the local definition. *)
- val map_value : (Constr.t -> Constr.t) -> t -> t
+ val map_value : ('c -> 'c) -> ('c, 't) pt -> ('c, 't) pt
(** Map the type of the name bound by a given declaration. *)
- val map_type : (Constr.t -> Constr.t) -> t -> t
+ val map_type : ('t -> 't) -> ('c, 't) pt -> ('c, 't) pt
(** Map all terms in a given declaration. *)
- val map_constr : (Constr.t -> Constr.t) -> t -> t
+ val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
(** Perform a given action on all terms in a given declaration. *)
- val iter_constr : (Constr.t -> unit) -> t -> unit
+ val iter_constr : ('c -> unit) -> ('c, 'c) pt -> unit
(** Reduce all terms in a given declaration to a single value. *)
- val fold_constr : (Constr.t -> 'a -> 'a) -> t -> 'a -> 'a
+ val fold_constr : ('c -> 'a -> 'a) -> ('c, 'c) pt -> 'a -> 'a
- val to_tuple : t -> Id.t * Constr.t option * Constr.t
- val of_tuple : Id.t * Constr.t option * Constr.t -> t
+ val to_tuple : ('c, 't) pt -> Id.t * 'c option * 't
+ val of_tuple : Id.t * 'c option * 't -> ('c, 't) pt
(** Convert [Rel.Declaration.t] value to the corresponding [Named.Declaration.t] value.
The function provided as the first parameter determines how to translate "names" to "ids". *)
- val of_rel_decl : (Name.t -> Id.t) -> Rel.Declaration.t -> t
+ val of_rel_decl : (Name.t -> Id.t) -> ('c, 't) Rel.Declaration.pt -> ('c, 't) pt
(** Convert [Named.Declaration.t] value to the corresponding [Rel.Declaration.t] value. *)
(* TODO: Move this function to [Rel.Declaration] module and rename it to [of_named]. *)
- val to_rel_decl : t -> Rel.Declaration.t
+ val to_rel_decl : ('c, 't) pt -> ('c, 't) Rel.Declaration.pt
end
(** Rel-context is represented as a list of declarations.
Inner-most declarations are at the beginning of the list.
Outer-most declarations are at the end of the list. *)
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
(** empty named-context *)
- val empty : t
+ val empty : ('c, 't) pt
(** Return a new rel-context enriched by with a given inner-most declaration. *)
- val add : Declaration.t -> t -> t
+ val add : ('c, 't) Declaration.pt -> ('c, 't) pt -> ('c, 't) pt
(** Return the number of {e local declarations} in a given named-context. *)
- val length : t -> int
+ val length : ('c, 't) pt -> int
(** Return a declaration designated by an identifier of the variable bound in that declaration.
@raise Not_found if the designated identifier is not bound in a given named-context. *)
- val lookup : Id.t -> t -> Declaration.t
+ val lookup : Id.t -> ('c, 't) pt -> ('c, 't) Declaration.pt
(** Check whether given two rel-contexts are equal. *)
- val equal : t -> t -> bool
+ val equal : ('c -> 'c -> bool) -> ('c, 'c) pt -> ('c, 'c) pt -> bool
(** Map all terms in a given named-context. *)
- val map : (Constr.t -> Constr.t) -> t -> t
+ val map : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
(** Perform a given action on every declaration in a given named-context. *)
- val iter : (Constr.t -> unit) -> t -> unit
+ val iter : ('c -> unit) -> ('c, 'c) pt -> unit
(** Reduce all terms in a given named-context to a single value.
Innermost declarations are processed first. *)
- val fold_inside : ('a -> Declaration.t -> 'a) -> init:'a -> t -> 'a
+ val fold_inside : ('a -> ('c, 't) Declaration.pt -> 'a) -> init:'a -> ('c, 't) pt -> 'a
(** Reduce all terms in a given named-context to a single value.
Outermost declarations are processed first. *)
- val fold_outside : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
+ val fold_outside : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a
(** Return the set of all identifiers bound in a given named-context. *)
- val to_vars : t -> Id.Set.t
+ val to_vars : ('c, 't) pt -> Id.Set.t
(** [instance_from_named_context Ω] builds an instance [args] such
that [Ω ⊢ args:Ω] where [Ω] is a named context and with the local
definitions of [Ω] skipped. Example: for [id1:T,id2:=c,id3:U], it
gives [Var id1, Var id3]. All [idj] are supposed distinct. *)
- val to_instance : t -> Constr.t list
+ val to_instance : (Id.t -> 'r) -> ('c, 't) pt -> 'r list
end
module Compacted :
sig
module Declaration :
sig
- type t =
- | LocalAssum of Id.t list * Constr.t
- | LocalDef of Id.t list * Constr.t * Constr.t
+ type ('constr, 'types) pt =
+ | LocalAssum of Id.t list * 'types
+ | LocalDef of Id.t list * 'constr * 'types
+
+ type t = (Constr.constr, Constr.types) pt
- val map_constr : (Constr.t -> Constr.t) -> t -> t
- val of_named_decl : Named.Declaration.t -> t
- val to_named_context : t -> Named.t
+ val map_constr : ('c -> 'c) -> ('c, 'c) pt -> ('c, 'c) pt
+ val of_named_decl : ('c, 't) Named.Declaration.pt -> ('c, 't) pt
+ val to_named_context : ('c, 't) pt -> ('c, 't) Named.pt
end
+ type ('constr, 'types) pt = ('constr, 'types) Declaration.pt list
type t = Declaration.t list
- val fold : (Declaration.t -> 'a -> 'a) -> t -> init:'a -> 'a
+ val fold : (('c, 't) Declaration.pt -> 'a -> 'a) -> ('c, 't) pt -> init:'a -> 'a
end
diff --git a/kernel/environ.ml b/kernel/environ.ml
index 4a543f195..9986f439a 100644
--- a/kernel/environ.ml
+++ b/kernel/environ.ml
@@ -120,7 +120,7 @@ let lookup_named = lookup_named
let lookup_named_val id ctxt = fst (Id.Map.find id ctxt.env_named_map)
let eq_named_context_val c1 c2 =
- c1 == c2 || Context.Named.equal (named_context_of_val c1) (named_context_of_val c2)
+ c1 == c2 || Context.Named.equal Constr.equal (named_context_of_val c1) (named_context_of_val c2)
(* A local const is evaluable if it is defined *)
@@ -469,9 +469,11 @@ let lookup_modtype mp env =
(*s Judgments. *)
-type unsafe_judgment = {
- uj_val : constr;
- uj_type : types }
+type ('constr, 'types) punsafe_judgment = {
+ uj_val : 'constr;
+ uj_type : 'types }
+
+type unsafe_judgment = (constr, types) punsafe_judgment
let make_judge v tj =
{ uj_val = v;
@@ -480,10 +482,12 @@ let make_judge v tj =
let j_val j = j.uj_val
let j_type j = j.uj_type
-type unsafe_type_judgment = {
- utj_val : constr;
+type 'types punsafe_type_judgment = {
+ utj_val : 'types;
utj_type : sorts }
+type unsafe_type_judgment = types punsafe_type_judgment
+
(*s Compilation of global declaration *)
let compile_constant_body = Cbytegen.compile_constant_body false
diff --git a/kernel/environ.mli b/kernel/environ.mli
index ea570cb4a..b7431dbe5 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -238,18 +238,21 @@ val keep_hyps : env -> Id.Set.t -> Context.Named.t
actually only a datatype to store a term with its type and the type of its
type. *)
-type unsafe_judgment = {
- uj_val : constr;
- uj_type : types }
+type ('constr, 'types) punsafe_judgment = {
+ uj_val : 'constr;
+ uj_type : 'types }
-val make_judge : constr -> types -> unsafe_judgment
-val j_val : unsafe_judgment -> constr
-val j_type : unsafe_judgment -> types
+type unsafe_judgment = (constr, types) punsafe_judgment
-type unsafe_type_judgment = {
- utj_val : constr;
+val make_judge : 'constr -> 'types -> ('constr, 'types) punsafe_judgment
+val j_val : ('constr, 'types) punsafe_judgment -> 'constr
+val j_type : ('constr, 'types) punsafe_judgment -> 'types
+
+type 'types punsafe_type_judgment = {
+ utj_val : 'types;
utj_type : sorts }
+type unsafe_type_judgment = types punsafe_type_judgment
(** {6 Compilation of global declaration } *)
diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml
index de97268b3..2ff419338 100644
--- a/kernel/indtypes.ml
+++ b/kernel/indtypes.ml
@@ -734,7 +734,7 @@ let compute_projections ((kn, _ as ind), u as indu) n x nparamargs params
matching with a parameter context. *)
let indty, paramsletsubst =
(* [ty] = [Ind inst] is typed in context [params] *)
- let inst = Context.Rel.to_extended_vect 0 paramslet in
+ let inst = Context.Rel.to_extended_vect mkRel 0 paramslet in
let ty = mkApp (mkIndU indu, inst) in
(* [Ind inst] is typed in context [params-wo-let] *)
let inst' = rel_list 0 nparamargs in
diff --git a/kernel/inductive.ml b/kernel/inductive.ml
index 3c4c2796e..d8d365c34 100644
--- a/kernel/inductive.ml
+++ b/kernel/inductive.ml
@@ -297,7 +297,7 @@ let build_dependent_inductive ind (_,mip) params =
applist
(mkIndU ind,
List.map (lift mip.mind_nrealdecls) params
- @ Context.Rel.to_extended_list 0 realargs)
+ @ Context.Rel.to_extended_list mkRel 0 realargs)
(* This exception is local *)
exception LocalArity of (sorts_family * sorts_family * arity_error) option
@@ -355,7 +355,7 @@ let build_branches_type (ind,u) (_,mip as specif) params p =
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 0 cstrsign)) in
+ let dep_cstr = 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
diff --git a/kernel/term.ml b/kernel/term.ml
index 62c161be4..e5a681375 100644
--- a/kernel/term.ml
+++ b/kernel/term.ml
@@ -71,20 +71,21 @@ type pconstant = constant puniverses
type pinductive = inductive puniverses
type pconstructor = constructor puniverses
-type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
+ ('constr, 'types, 'sort, 'univs) Constr.kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of sorts
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of pconstant
- | Ind of pinductive
- | Construct of pconstructor
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
@@ -179,8 +180,6 @@ let destMeta c = match kind_of_term c with
| _ -> raise DestKO
let isMeta c = match kind_of_term c with Meta _ -> true | _ -> false
-let isMetaOf mv c =
- match kind_of_term c with Meta mv' -> Int.equal mv mv' | _ -> false
(* Destructs a variable *)
let destVar c = match kind_of_term c with
diff --git a/kernel/term.mli b/kernel/term.mli
index a8d9dfbff..a9bb67705 100644
--- a/kernel/term.mli
+++ b/kernel/term.mli
@@ -68,20 +68,21 @@ type ('constr, 'types) prec_declaration =
type ('constr, 'types) pfixpoint = ('constr, 'types) Constr.pfixpoint
type ('constr, 'types) pcofixpoint = ('constr, 'types) Constr.pcofixpoint
-type ('constr, 'types) kind_of_term = ('constr, 'types) Constr.kind_of_term =
+type ('constr, 'types, 'sort, 'univs) kind_of_term =
+ ('constr, 'types, 'sort, 'univs) Constr.kind_of_term =
| Rel of int
| Var of Id.t
| Meta of metavariable
| Evar of 'constr pexistential
- | Sort of sorts
+ | Sort of 'sort
| Cast of 'constr * cast_kind * 'types
| Prod of Name.t * 'types * 'types
| Lambda of Name.t * 'types * 'constr
| LetIn of Name.t * 'constr * 'types * 'constr
| App of 'constr * 'constr array
- | Const of constant puniverses
- | Ind of inductive puniverses
- | Construct of constructor puniverses
+ | Const of (constant * 'univs)
+ | Ind of (inductive * 'univs)
+ | Construct of (constructor * 'univs)
| Case of case_info * 'constr * 'constr * 'constr array
| Fix of ('constr, 'types) pfixpoint
| CoFix of ('constr, 'types) pcofixpoint
@@ -98,7 +99,6 @@ val isVarId : Id.t -> constr -> bool
val isInd : constr -> bool
val isEvar : constr -> bool
val isMeta : constr -> bool
-val isMetaOf : metavariable -> constr -> bool
val isEvar_or_Meta : constr -> bool
val isSort : constr -> bool
val isCast : constr -> bool
@@ -444,7 +444,7 @@ val leq_constr_univs : constr UGraph.check_function
application grouping and ignoring universe instances. *)
val eq_constr_nounivs : constr -> constr -> bool
-val kind_of_term : constr -> (constr, types) kind_of_term
+val kind_of_term : constr -> (constr, types, Sorts.t, Univ.Instance.t) kind_of_term
(** Alias for [Constr.kind] *)
val constr_ord : constr -> constr -> int
diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml
index 5071f0ad5..5e1763815 100644
--- a/kernel/type_errors.ml
+++ b/kernel/type_errors.ml
@@ -13,52 +13,56 @@ open Reduction
(* Type errors. *)
-type guard_error =
+type 'constr pguard_error =
(* Fixpoints *)
| NotEnoughAbstractionInFixBody
- | RecursionNotOnInductiveType of constr
- | RecursionOnIllegalTerm of int * (env * constr) * int list * int list
+ | RecursionNotOnInductiveType of 'constr
+ | RecursionOnIllegalTerm of int * (env * 'constr) * int list * int list
| NotEnoughArgumentsForFixCall of int
(* CoFixpoints *)
- | CodomainNotInductiveType of constr
+ | CodomainNotInductiveType of 'constr
| NestedRecursiveOccurrences
- | UnguardedRecursiveCall of constr
- | RecCallInTypeOfAbstraction of constr
- | RecCallInNonRecArgOfConstructor of constr
- | RecCallInTypeOfDef of constr
- | RecCallInCaseFun of constr
- | RecCallInCaseArg of constr
- | RecCallInCasePred of constr
- | NotGuardedForm of constr
- | ReturnPredicateNotCoInductive of constr
+ | UnguardedRecursiveCall of 'constr
+ | RecCallInTypeOfAbstraction of 'constr
+ | RecCallInNonRecArgOfConstructor of 'constr
+ | RecCallInTypeOfDef of 'constr
+ | RecCallInCaseFun of 'constr
+ | RecCallInCaseArg of 'constr
+ | RecCallInCasePred of 'constr
+ | NotGuardedForm of 'constr
+ | ReturnPredicateNotCoInductive of 'constr
+
+type guard_error = constr pguard_error
type arity_error =
| NonInformativeToInformative
| StrongEliminationOnNonSmallType
| WrongArity
-type type_error =
+type ('constr, 'types) ptype_error =
| UnboundRel of int
| UnboundVar of variable
- | NotAType of unsafe_judgment
- | BadAssumption of unsafe_judgment
- | ReferenceVariables of identifier * constr
- | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
+ | NotAType of ('constr, 'types) punsafe_judgment
+ | BadAssumption of ('constr, 'types) punsafe_judgment
+ | ReferenceVariables of identifier * 'constr
+ | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment
* (sorts_family * sorts_family * arity_error) option
- | CaseNotInductive of unsafe_judgment
+ | CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
- | NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * pconstructor * constr * constr
- | Generalization of (Name.t * types) * unsafe_judgment
- | ActualType of unsafe_judgment * types
+ | NumberBranches of ('constr, 'types) punsafe_judgment * int
+ | IllFormedBranch of 'constr * pconstructor * 'constr * 'constr
+ | Generalization of (Name.t * 'types) * ('constr, 'types) punsafe_judgment
+ | ActualType of ('constr, 'types) punsafe_judgment * 'types
| CantApplyBadType of
- (int * constr * constr) * unsafe_judgment * unsafe_judgment array
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
- | IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array
+ (int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array
| IllTypedRecBody of
- int * Name.t array * unsafe_judgment array * types array
+ int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array
| UnsatisfiedConstraints of Univ.constraints
+type type_error = (constr, types) ptype_error
+
exception TypeError of env * type_error
let nfj env {uj_val=c;uj_type=ct} =
diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli
index 0c3a952b8..bd6032716 100644
--- a/kernel/type_errors.mli
+++ b/kernel/type_errors.mli
@@ -14,52 +14,56 @@ open Environ
(*i Rem: NotEnoughAbstractionInFixBody should only occur with "/i" Fix
notation i*)
-type guard_error =
+type 'constr pguard_error =
(** Fixpoints *)
| NotEnoughAbstractionInFixBody
- | RecursionNotOnInductiveType of constr
- | RecursionOnIllegalTerm of int * (env * constr) * int list * int list
+ | RecursionNotOnInductiveType of 'constr
+ | RecursionOnIllegalTerm of int * (env * 'constr) * int list * int list
| NotEnoughArgumentsForFixCall of int
(** CoFixpoints *)
- | CodomainNotInductiveType of constr
+ | CodomainNotInductiveType of 'constr
| NestedRecursiveOccurrences
- | UnguardedRecursiveCall of constr
- | RecCallInTypeOfAbstraction of constr
- | RecCallInNonRecArgOfConstructor of constr
- | RecCallInTypeOfDef of constr
- | RecCallInCaseFun of constr
- | RecCallInCaseArg of constr
- | RecCallInCasePred of constr
- | NotGuardedForm of constr
- | ReturnPredicateNotCoInductive of constr
+ | UnguardedRecursiveCall of 'constr
+ | RecCallInTypeOfAbstraction of 'constr
+ | RecCallInNonRecArgOfConstructor of 'constr
+ | RecCallInTypeOfDef of 'constr
+ | RecCallInCaseFun of 'constr
+ | RecCallInCaseArg of 'constr
+ | RecCallInCasePred of 'constr
+ | NotGuardedForm of 'constr
+ | ReturnPredicateNotCoInductive of 'constr
+
+type guard_error = constr pguard_error
type arity_error =
| NonInformativeToInformative
| StrongEliminationOnNonSmallType
| WrongArity
-type type_error =
+type ('constr, 'types) ptype_error =
| UnboundRel of int
| UnboundVar of variable
- | NotAType of unsafe_judgment
- | BadAssumption of unsafe_judgment
- | ReferenceVariables of identifier * constr
- | ElimArity of pinductive * sorts_family list * constr * unsafe_judgment
+ | NotAType of ('constr, 'types) punsafe_judgment
+ | BadAssumption of ('constr, 'types) punsafe_judgment
+ | ReferenceVariables of identifier * 'constr
+ | ElimArity of pinductive * sorts_family list * 'constr * ('constr, 'types) punsafe_judgment
* (sorts_family * sorts_family * arity_error) option
- | CaseNotInductive of unsafe_judgment
+ | CaseNotInductive of ('constr, 'types) punsafe_judgment
| WrongCaseInfo of pinductive * case_info
- | NumberBranches of unsafe_judgment * int
- | IllFormedBranch of constr * pconstructor * constr * constr
- | Generalization of (Name.t * types) * unsafe_judgment
- | ActualType of unsafe_judgment * types
+ | NumberBranches of ('constr, 'types) punsafe_judgment * int
+ | IllFormedBranch of 'constr * pconstructor * 'constr * 'constr
+ | Generalization of (Name.t * 'types) * ('constr, 'types) punsafe_judgment
+ | ActualType of ('constr, 'types) punsafe_judgment * 'types
| CantApplyBadType of
- (int * constr * constr) * unsafe_judgment * unsafe_judgment array
- | CantApplyNonFunctional of unsafe_judgment * unsafe_judgment array
- | IllFormedRecBody of guard_error * Name.t array * int * env * unsafe_judgment array
+ (int * 'constr * 'constr) * ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | CantApplyNonFunctional of ('constr, 'types) punsafe_judgment * ('constr, 'types) punsafe_judgment array
+ | IllFormedRecBody of 'constr pguard_error * Name.t array * int * env * ('constr, 'types) punsafe_judgment array
| IllTypedRecBody of
- int * Name.t array * unsafe_judgment array * types array
+ int * Name.t array * ('constr, 'types) punsafe_judgment array * 'types array
| UnsatisfiedConstraints of Univ.constraints
+type type_error = (constr, types) ptype_error
+
exception TypeError of env * type_error
val error_unbound_rel : env -> int -> 'a
diff --git a/kernel/vars.mli b/kernel/vars.mli
index f7535e6d8..adeac422e 100644
--- a/kernel/vars.mli
+++ b/kernel/vars.mli
@@ -75,7 +75,7 @@ val adjust_subst_to_rel_context : Context.Rel.t -> constr list -> constr list
(** Take an index in an instance of a context and returns its index wrt to
the full context (e.g. 2 in [x:A;y:=b;z:C] is 3, i.e. a reference to z) *)
-val adjust_rel_to_rel_context : Context.Rel.t -> int -> int
+val adjust_rel_to_rel_context : ('a, 'b) Context.Rel.pt -> int -> int
(** [substnl [a₁;...;an] k c] substitutes in parallel [a₁],...,[an]
for respectively [Rel(k+1)],...,[Rel(k+n)] in [c]; it relocates
diff --git a/library/impargs.ml b/library/impargs.ml
index 836568b89..a63264b66 100644
--- a/library/impargs.ml
+++ b/library/impargs.ml
@@ -228,10 +228,12 @@ let rec is_rigid_head t = match kind_of_term t with
(* calcule la liste des arguments implicites *)
-let find_displayed_name_in all avoid na (_,b as envnames_b) =
+let find_displayed_name_in all avoid na (env, b) =
+ let b = EConstr.of_constr b in
+ let envnames_b = (env, b) in
let flag = RenamingElsewhereFor envnames_b in
- if all then compute_and_force_displayed_name_in flag avoid na b
- else compute_displayed_name_in flag avoid na b
+ if all then compute_and_force_displayed_name_in Evd.empty flag avoid na b
+ else compute_displayed_name_in Evd.empty flag avoid na b
let compute_implicits_gen strict strongly_strict revpat contextual all env t =
let rigid = ref true in
diff --git a/library/keys.ml b/library/keys.ml
index 057dc3b65..c9e325ee5 100644
--- a/library/keys.ml
+++ b/library/keys.ml
@@ -114,11 +114,11 @@ let inKeys : key_obj -> obj =
let declare_equiv_keys ref ref' =
Lib.add_anonymous_leaf (inKeys (ref,ref'))
-let constr_key c =
+let constr_key kind c =
let open Globnames in
try
let rec aux k =
- match kind_of_term k with
+ match kind k with
| Const (c, _) -> KGlob (ConstRef c)
| Ind (i, u) -> KGlob (IndRef i)
| Construct (c,u) -> KGlob (ConstructRef c)
diff --git a/library/keys.mli b/library/keys.mli
index 69668590d..6abac4de4 100644
--- a/library/keys.mli
+++ b/library/keys.mli
@@ -16,7 +16,7 @@ val declare_equiv_keys : key -> key -> unit
val equiv_keys : key -> key -> bool
(** Check equivalence of keys. *)
-val constr_key : Term.constr -> key option
+val constr_key : ('a -> ('a, 't, 'u, 'i) Constr.kind_of_term) -> 'a -> key option
(** Compute the head key of a term. *)
val pr_keys : (global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds
diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml
index 2c5b108e5..a0b04ce3b 100644
--- a/plugins/btauto/refl_btauto.ml
+++ b/plugins/btauto/refl_btauto.ml
@@ -14,8 +14,8 @@ 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 (c : Term.constr) =
- Term.kind_of_term (Termops.strip_outer_cast c)
+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 lapp c v = Term.mkApp (Lazy.force c, v)
@@ -105,7 +105,7 @@ module Bool = struct
| Negb of t
| Ifb of t * t * t
- let quote (env : Env.t) (c : Term.constr) : t =
+ let quote (env : Env.t) sigma (c : Term.constr) : t =
let trueb = Lazy.force trueb in
let falseb = Lazy.force falseb in
let andb = Lazy.force andb in
@@ -113,7 +113,7 @@ module Bool = struct
let xorb = Lazy.force xorb in
let negb = Lazy.force negb in
- let rec aux c = match decomp_term c with
+ let rec aux c = match decomp_term sigma c with
| Term.App (head, args) ->
if head === andb && Array.length args = 2 then
Andb (aux args.(0), aux args.(1))
@@ -179,9 +179,11 @@ module Btauto = struct
let print_counterexample p env gl =
let var = lapp witness [|p|] in
+ let var = EConstr.of_constr var in
(* Compute an assignment that dissatisfies the goal *)
let _, var = Tacmach.pf_reduction_of_red_expr gl (Genredexpr.CbvVm None) var in
- let rec to_list l = match decomp_term l with
+ let var = EConstr.Unsafe.to_constr var in
+ let rec to_list l = match decomp_term (Tacmach.project gl) l with
| Term.App (c, _)
when c === (Lazy.force CoqList._nil) -> []
| Term.App (c, [|_; h; t|])
@@ -220,7 +222,8 @@ module Btauto = struct
Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let eq = Lazy.force eq in
- let t = decomp_term concl in
+ let concl = EConstr.Unsafe.to_constr concl in
+ let t = decomp_term (Tacmach.New.project gl) concl in
match t with
| Term.App (c, [|typ; p; _|]) when c === eq ->
(* should be an equality [@eq poly ?p (Cst false)] *)
@@ -234,22 +237,25 @@ module Btauto = struct
let tac =
Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
+ let concl = EConstr.Unsafe.to_constr concl in
+ let sigma = Tacmach.New.project gl in
let eq = Lazy.force eq in
let bool = Lazy.force Bool.typ in
- let t = decomp_term concl in
+ let t = decomp_term sigma concl in
match t with
| Term.App (c, [|typ; tl; tr|])
when typ === bool && c === eq ->
let env = Env.empty () in
- let fl = Bool.quote env tl in
- let fr = Bool.quote env tr in
+ let fl = Bool.quote env sigma tl in
+ let fr = Bool.quote env sigma tr in
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 = EConstr.of_constr changed_gl in
Tacticals.New.tclTHENLIST [
Tactics.change_concl changed_gl;
- Tactics.apply (Lazy.force soundness);
+ Tactics.apply (EConstr.of_constr (Lazy.force soundness));
Tactics.normalise_vm_in_concl;
try_unification env
]
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 7347c3c2c..aa71a4565 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -452,8 +452,9 @@ and applist_projection c l =
applistc (mkProj (p, hd)) tl)
| _ -> applistc c l
-let rec canonize_name c =
- let func = canonize_name in
+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
| Const (kn,u) ->
let canon_const = constant_of_kn (canonical_con kn) in
@@ -497,10 +498,10 @@ let rec inst_pattern subst = function
args t
let pr_idx_term uf i = str "[" ++ int i ++ str ":=" ++
- Termops.print_constr (constr_of_term (term uf i)) ++ str "]"
+ Termops.print_constr (EConstr.of_constr (constr_of_term (term uf i))) ++ str "]"
let pr_term t = str "[" ++
- Termops.print_constr (constr_of_term t) ++ str "]"
+ Termops.print_constr (EConstr.of_constr (constr_of_term t)) ++ str "]"
let rec add_term state t=
let uf=state.uf in
@@ -508,8 +509,8 @@ let rec add_term state t=
Not_found ->
let b=next uf in
let trm = constr_of_term t in
- let typ = pf_unsafe_type_of state.gls trm in
- let typ = canonize_name typ in
+ let typ = pf_unsafe_type_of state.gls (EConstr.of_constr trm) in
+ let typ = canonize_name (project state.gls) typ in
let new_node=
match t with
Symb _ | Product (_,_) ->
@@ -615,7 +616,7 @@ let add_inst state (inst,int_subst) =
begin
debug (fun () ->
(str "Adding new equality, depth="++ int state.rew_depth) ++ fnl () ++
- (str " [" ++ Termops.print_constr prf ++ str " : " ++
+ (str " [" ++ Termops.print_constr (EConstr.of_constr prf) ++ str " : " ++
pr_term s ++ str " == " ++ pr_term t ++ str "]"));
add_equality state prf s t
end
@@ -623,7 +624,7 @@ let add_inst state (inst,int_subst) =
begin
debug (fun () ->
(str "Adding new disequality, depth="++ int state.rew_depth) ++ fnl () ++
- (str " [" ++ Termops.print_constr prf ++ str " : " ++
+ (str " [" ++ Termops.print_constr (EConstr.of_constr prf) ++ str " : " ++
pr_term s ++ str " <> " ++ pr_term t ++ str "]"));
add_disequality state (Hyp prf) s t
end
@@ -832,7 +833,8 @@ let complete_one_class state i=
let id = new_state_var etyp state in
app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in
let _c = pf_unsafe_type_of state.gls
- (constr_of_term (term state.uf pac.cnode)) in
+ (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in
+ let _c = EConstr.Unsafe.to_constr _c in
let _args =
List.map (fun i -> constr_of_term (term state.uf i))
pac.args in
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index b5ca2f50f..2d9dec095 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -13,6 +13,7 @@ open Names
open Inductiveops
open Declarations
open Term
+open EConstr
open Vars
open Tacmach
open Tactics
@@ -39,13 +40,11 @@ let _False = reference ["Init";"Logic"] "False"
let _True = reference ["Init";"Logic"] "True"
let _I = reference ["Init";"Logic"] "I"
-let whd env=
- let infos=CClosure.create_clos_infos CClosure.betaiotazeta env in
- (fun t -> CClosure.whd_val infos (CClosure.inject t))
+let whd env sigma t =
+ Reductionops.clos_whd_flags CClosure.betaiotazeta env sigma t
-let whd_delta env=
- let infos=CClosure.create_clos_infos CClosure.all env in
- (fun t -> CClosure.whd_val infos (CClosure.inject t))
+let whd_delta env sigma t =
+ Reductionops.clos_whd_flags CClosure.all env sigma t
(* decompose member of equality in an applicative format *)
@@ -53,12 +52,12 @@ let whd_delta env=
let sf_of env sigma c = e_sort_of env (ref sigma) c
let rec decompose_term env sigma t=
- match kind_of_term (whd env t) with
+ match EConstr.kind sigma (whd env sigma t) with
App (f,args)->
let tf=decompose_term env sigma f in
let targs=Array.map (decompose_term env sigma) args in
Array.fold_left (fun s t->Appli (s,t)) tf targs
- | Prod (_,a,_b) when not (Termops.dependent (mkRel 1) _b) ->
+ | Prod (_,a,_b) when noccurn sigma 1 _b ->
let b = Termops.pop _b in
let sort_b = sf_of env sigma b in
let sort_a = sf_of env sigma a in
@@ -67,6 +66,7 @@ let rec decompose_term env sigma t=
decompose_term env sigma b)
| Construct c ->
let (((mind,i_ind),i_con),u)= c in
+ let u = EInstance.kind sigma u in
let canon_mind = mind_of_kn (canonical_mind mind) in
let canon_ind = canon_mind,i_ind in
let (oib,_)=Global.lookup_inductive (canon_ind) in
@@ -76,28 +76,30 @@ let rec decompose_term env sigma t=
ci_nhyps=nargs-oib.mind_nparams}
| Ind c ->
let (mind,i_ind),u = c in
+ let u = EInstance.kind sigma u in
let canon_mind = mind_of_kn (canonical_mind mind) in
- let canon_ind = canon_mind,i_ind in (Symb (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_of_kn (canonical_con c) in
- (Symb (mkConstU (canon_const,u)))
+ (Symb (Constr.mkConstU (canon_const,u)))
| Proj (p, c) ->
let canon_const kn = constant_of_kn (canonical_con kn) in
let p' = Projection.map canon_const p in
- (Appli (Symb (mkConst (Projection.constant p')), decompose_term env sigma c))
+ (Appli (Symb (Constr.mkConst (Projection.constant p')), decompose_term env sigma c))
| _ ->
- let t = Termops.strip_outer_cast t in
- if closed0 t then Symb t else raise Not_found
+ let t = Termops.strip_outer_cast sigma t in
+ if closed0 sigma t then Symb (EConstr.to_constr sigma t) else raise Not_found
(* decompose equality in members and type *)
-open Globnames
+open Termops
let atom_of_constr env sigma term =
- let wh = (whd_delta env term) in
- let kot = kind_of_term wh in
+ let wh = whd_delta env sigma term in
+ let kot = EConstr.kind sigma wh in
match kot with
App (f,args)->
- if is_global (Lazy.force _eq) f && Int.equal (Array.length args) 3
+ if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3
then `Eq (args.(0),
decompose_term env sigma args.(1),
decompose_term env sigma args.(2))
@@ -105,14 +107,14 @@ let atom_of_constr env sigma term =
| _ -> `Other (decompose_term env sigma term)
let rec pattern_of_constr env sigma c =
- match kind_of_term (whd env c) with
+ match EConstr.kind sigma (whd env sigma c) with
App (f,args)->
let pf = decompose_term env sigma f in
let pargs,lrels = List.split
(Array.map_to_list (pattern_of_constr env sigma) args) in
PApp (pf,List.rev pargs),
List.fold_left Int.Set.union Int.Set.empty lrels
- | Prod (_,a,_b) when not (Termops.dependent (mkRel 1) _b) ->
+ | Prod (_,a,_b) when noccurn sigma 1 _b ->
let b = Termops.pop _b in
let pa,sa = pattern_of_constr env sigma a in
let pb,sb = pattern_of_constr env sigma b in
@@ -131,19 +133,19 @@ let non_trivial = function
let patterns_of_constr env sigma nrels term=
let f,args=
- try destApp (whd_delta env term) with DestKO -> raise Not_found in
- if is_global (Lazy.force _eq) f && Int.equal (Array.length args) 3
+ try destApp sigma (whd_delta env sigma term) with DestKO -> raise Not_found in
+ if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3
then
let patt1,rels1 = pattern_of_constr env sigma args.(1)
and patt2,rels2 = pattern_of_constr env sigma args.(2) in
let valid1 =
if not (Int.equal (Int.Set.cardinal rels1) nrels) then Creates_variables
else if non_trivial patt1 then Normal
- else Trivial args.(0)
+ else Trivial (EConstr.to_constr sigma args.(0))
and valid2 =
if not (Int.equal (Int.Set.cardinal rels2) nrels) then Creates_variables
else if non_trivial patt2 then Normal
- else Trivial args.(0) in
+ else Trivial (EConstr.to_constr sigma args.(0)) in
if valid1 != Creates_variables
|| valid2 != Creates_variables then
nrels,valid1,patt1,valid2,patt2
@@ -151,28 +153,28 @@ let patterns_of_constr env sigma nrels term=
else raise Not_found
let rec quantified_atom_of_constr env sigma nrels term =
- match kind_of_term (whd_delta env term) with
+ match EConstr.kind sigma (whd_delta env sigma term) with
Prod (id,atom,ff) ->
- if is_global (Lazy.force _False) ff then
+ if is_global sigma (Lazy.force _False) ff then
let patts=patterns_of_constr env sigma nrels atom in
`Nrule patts
else
- quantified_atom_of_constr (Environ.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff
+ quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff
| _ ->
let patts=patterns_of_constr env sigma nrels term in
`Rule patts
let litteral_of_constr env sigma term=
- match kind_of_term (whd_delta env term) with
+ match EConstr.kind sigma (whd_delta env sigma term) with
| Prod (id,atom,ff) ->
- if is_global (Lazy.force _False) ff then
+ if is_global sigma (Lazy.force _False) ff then
match (atom_of_constr env sigma atom) with
`Eq(t,a,b) -> `Neq(t,a,b)
| `Other(p) -> `Nother(p)
else
begin
try
- quantified_atom_of_constr (Environ.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff
+ quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff
with Not_found ->
`Other (decompose_term env sigma term)
end
@@ -183,9 +185,10 @@ let litteral_of_constr env sigma term=
(* store all equalities from the context *)
let make_prb gls depth additionnal_terms =
+ let open Tacmach.New in
let env=pf_env gls in
- let sigma=sig_sig gls in
- let state = empty depth gls in
+ let sigma=project gls in
+ let state = empty depth {it = Proofview.Goal.goal (Proofview.Goal.assume gls); sigma } in
let pos_hyps = ref [] in
let neg_hyps =ref [] in
List.iter
@@ -196,7 +199,7 @@ let make_prb gls depth additionnal_terms =
(fun decl ->
let id = NamedDecl.get_id decl in
begin
- let cid=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
@@ -214,9 +217,9 @@ let make_prb gls depth additionnal_terms =
neg_hyps:=(cid,nh):: !neg_hyps
| `Rule patts -> add_quant state id true patts
| `Nrule patts -> add_quant state id false patts
- end) (Environ.named_context_of_val (Goal.V82.nf_hyps gls.sigma gls.it));
+ end) (Proofview.Goal.hyps gls);
begin
- match atom_of_constr env sigma (Evarutil.nf_evar sigma (pf_concl gls)) with
+ match atom_of_constr env sigma (pf_concl gls) with
`Eq (t,a,b) -> add_disequality state Goal a b
| `Other g ->
List.iter
@@ -228,6 +231,7 @@ let make_prb gls depth additionnal_terms =
(* indhyps builds the array of arrays of constructor hyps for (ind largs) *)
let build_projection intype (cstr:pconstructor) special default gls=
+ let open Tacmach.New in
let ci= (snd(fst cstr)) in
let body=Equality.build_selector (pf_env gls) (project gls) ci (mkRel 1) intype special default in
let id=pf_get_new_id (Id.of_string "t") gls in
@@ -244,6 +248,7 @@ let new_app_global f args k =
Tacticals.New.pf_constr_of_global (Lazy.force f) (fun fc -> k (mkApp (fc, args)))
let new_refine c = Proofview.V82.tactic (refine c)
+let refine c = refine c
let assert_before n c =
Proofview.Goal.enter { enter = begin fun gl ->
@@ -256,20 +261,23 @@ let refresh_type env evm ty =
(Some false) env evm ty
let refresh_universes ty k =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let evm = Tacmach.New.project gl in
let evm, ty = refresh_type env evm ty in
- Tacticals.New.tclTHEN (Proofview.V82.tactic (Refiner.tclEVARS evm)) (k ty)
+ Sigma.Unsafe.of_pair (k ty, evm)
end }
+let constr_of_term c = EConstr.of_constr (constr_of_term c)
+
let rec proof_tac p : unit Proofview.tactic =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let type_of t = Tacmach.New.pf_unsafe_type_of gl t in
try (* type_of can raise exceptions *)
match p.p_rule with
- Ax c -> exact_check c
+ Ax c -> exact_check (EConstr.of_constr c)
| SymAx c ->
+ let c = EConstr.of_constr c in
let l=constr_of_term p.p_lhs and
r=constr_of_term p.p_rhs in
refresh_universes (type_of l) (fun typ ->
@@ -293,7 +301,7 @@ let rec proof_tac p : unit Proofview.tactic =
refresh_universes (type_of tf1) (fun typf ->
refresh_universes (type_of tx1) (fun typx ->
refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx ->
- let id = Tacmach.New.of_old (fun gls -> pf_get_new_id (Id.of_string "f") gls) gl in
+ let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in
let appx1 = mkLambda(Name id,typf,mkApp(mkRel 1,[|tx1|])) in
let lemma1 = app_global _f_equal [|typf;typfx;appx1;tf1;tf2;_M 1|] in
let lemma2 = app_global _f_equal [|typx;typfx;tf2;tx1;tx2;_M 1|] in
@@ -319,7 +327,7 @@ let rec proof_tac p : unit Proofview.tactic =
refresh_universes (type_of ti) (fun intype ->
refresh_universes (type_of default) (fun outtype ->
let proj =
- Tacmach.New.of_old (build_projection intype cstr special default) gl
+ build_projection intype cstr special default gl
in
let injt=
app_global _f_equal [|intype;outtype;proj;ti;tj;_M 1|] in
@@ -328,9 +336,9 @@ let rec proof_tac p : unit Proofview.tactic =
end }
let refute_tac c t1 t2 p =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let tt1=constr_of_term t1 and tt2=constr_of_term t2 in
- let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in
+ let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in
let false_t=mkApp (c,[|mkVar hid|]) in
let k intype =
let neweq= new_app_global _eq [|intype;tt1;tt2|] in
@@ -344,12 +352,12 @@ let refine_exact_check c gl =
Tacticals.tclTHEN (Refiner.tclEVARS evm) (Proofview.V82.of_tactic (exact_check c)) gl
let convert_to_goal_tac c t1 t2 p =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let tt1=constr_of_term t1 and tt2=constr_of_term t2 in
let k sort =
let neweq= new_app_global _eq [|sort;tt1;tt2|] in
- let e = Tacmach.New.of_old (pf_get_new_id (Id.of_string "e")) gl in
- let x = Tacmach.New.of_old (pf_get_new_id (Id.of_string "X")) gl in
+ let e = Tacmach.New.pf_get_new_id (Id.of_string "e") gl in
+ let x = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in
let identity=mkLambda (Name x,sort,mkRel 1) in
let endt=app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in
Tacticals.New.tclTHENS (neweq (assert_before (Name e)))
@@ -358,9 +366,9 @@ let convert_to_goal_tac c t1 t2 p =
end }
let convert_to_hyp_tac c1 t1 c2 t2 p =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let tt2=constr_of_term t2 in
- let h = Tacmach.New.of_old (pf_get_new_id (Id.of_string "H")) gl in
+ let h = Tacmach.New.pf_get_new_id (Id.of_string "H") gl in
let false_t=mkApp (c2,[|mkVar h|]) in
Tacticals.New.tclTHENS (assert_before (Name h) tt2)
[convert_to_goal_tac c1 t1 t2 p;
@@ -368,20 +376,22 @@ let convert_to_hyp_tac c1 t1 c2 t2 p =
end }
let discriminate_tac (cstr,u as cstru) p =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let t1=constr_of_term p.p_lhs and t2=constr_of_term p.p_rhs in
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl gl in
- let xid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "X")) gl in
+ let xid = Tacmach.New.pf_get_new_id (Id.of_string "X") gl in
let identity = Universes.constr_of_global (Lazy.force _I) in
+ let identity = EConstr.of_constr identity in
let trivial = Universes.constr_of_global (Lazy.force _True) in
+ let trivial = EConstr.of_constr trivial in
let evm = Tacmach.New.project gl in
let evm, intype = refresh_type env evm (Tacmach.New.pf_unsafe_type_of gl t1) in
let evm, outtype = Evd.new_sort_variable Evd.univ_flexible evm in
let outtype = mkSort outtype in
let pred = mkLambda(Name xid,outtype,mkRel 1) in
- let hid = Tacmach.New.of_old (pf_get_new_id (Id.of_string "Heq")) gl in
- let proj = Tacmach.New.of_old (build_projection intype cstru trivial concl) gl in
+ let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in
+ let proj = build_projection intype cstru trivial concl gl in
let injt=app_global _f_equal
[|intype;outtype;proj;t1;t2;mkVar hid|] in
let endt k =
@@ -401,13 +411,15 @@ let build_term_to_complete uf meta pac =
let real_args = List.map (fun i -> constr_of_term (term uf i)) pac.args in
let dummy_args = List.rev (List.init pac.arity meta) in
let all_args = List.rev_append real_args dummy_args in
- applistc (mkConstructU cinfo.ci_constr) all_args
+ let (kn, u) = cinfo.ci_constr in
+ applist (mkConstructU (kn, EInstance.make u), all_args)
let cc_tactic depth additionnal_terms =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
Coqlib.check_required_library Coqlib.logic_module_name;
let _ = debug (fun () -> Pp.str "Reading subgoal ...") in
- let state = Tacmach.New.of_old (fun gls -> make_prb gls depth additionnal_terms) gl in
+ let state = make_prb gl depth additionnal_terms in
let _ = debug (fun () -> Pp.str "Problem built, solving ...") in
let sol = execute true state in
let _ = debug (fun () -> Pp.str "Computation completed.") in
@@ -439,7 +451,7 @@ let cc_tactic depth additionnal_terms =
str "\"congruence with (" ++
prlist_with_sep
(fun () -> str ")" ++ spc () ++ str "(")
- (Termops.print_constr_env env)
+ (Termops.print_constr_env env sigma)
terms_to_complete ++
str ")\","
end ++
@@ -450,10 +462,13 @@ let cc_tactic depth additionnal_terms =
let ta=term uf dis.lhs and tb=term uf dis.rhs in
match dis.rule with
Goal -> proof_tac p
- | Hyp id -> refute_tac id ta tb p
+ | Hyp id -> refute_tac (EConstr.of_constr id) ta tb p
| HeqG id ->
+ let id = EConstr.of_constr id in
convert_to_goal_tac id ta tb p
| HeqnH (ida,idb) ->
+ let ida = EConstr.of_constr ida in
+ let idb = EConstr.of_constr idb in
convert_to_hyp_tac ida ta idb tb p
end }
@@ -489,8 +504,9 @@ let mk_eq f c1 c2 k =
end })
let f_equal =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
let cut_eq c1 c2 =
try (* type_of can raise an exception *)
Tacticals.New.tclTHENS
@@ -499,9 +515,9 @@ let f_equal =
with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
in
Proofview.tclORELSE
- begin match kind_of_term concl with
- | App (r,[|_;t;t'|]) when Globnames.is_global (Lazy.force _eq) r ->
- begin match kind_of_term t, kind_of_term t' with
+ begin match EConstr.kind sigma concl with
+ | App (r,[|_;t;t'|]) when is_global sigma (Lazy.force _eq) r ->
+ begin match EConstr.kind sigma t, EConstr.kind sigma t' with
| App (f,v), App (f',v') when Int.equal (Array.length v) (Array.length v') ->
let rec cuts i =
if i < 0 then Tacticals.New.tclTRY (congruence_tac 1000 [])
@@ -512,7 +528,7 @@ let f_equal =
| _ -> Proofview.tclUNIT ()
end
begin function (e, info) -> match e with
- | Type_errors.TypeError _ -> Proofview.tclUNIT ()
+ | Pretype_errors.PretypeError _ | Type_errors.TypeError _ -> Proofview.tclUNIT ()
| e -> Proofview.tclZERO ~info e
end
end }
diff --git a/plugins/cc/cctac.mli b/plugins/cc/cctac.mli
index 7c1d9f1c0..de6eb982e 100644
--- a/plugins/cc/cctac.mli
+++ b/plugins/cc/cctac.mli
@@ -7,7 +7,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open EConstr
open Proof_type
val proof_tac: Ccproof.proof -> unit Proofview.tactic
diff --git a/plugins/derive/derive.ml b/plugins/derive/derive.ml
index e39d17b52..12d7f0660 100644
--- a/plugins/derive/derive.ml
+++ b/plugins/derive/derive.ml
@@ -28,12 +28,14 @@ let start_deriving f suchthat lemma =
(* spiwack: I don't know what the rigidity flag does, picked the one
that looked the most general. *)
let (sigma,f_type_sort) = Evd.new_sort_variable Evd.univ_flexible_alg sigma in
- let f_type_type = Term.mkSort f_type_sort in
+ let f_type_type = EConstr.mkSort f_type_sort in
(** create the initial goals for the proof: |- Type ; |- ?1 ; f:=?2 |- suchthat *)
let goals =
let open Proofview in
TCons ( env , sigma , f_type_type , (fun sigma f_type ->
TCons ( env , sigma , f_type , (fun sigma ef ->
+ let f_type = EConstr.Unsafe.to_constr f_type in
+ let ef = EConstr.Unsafe.to_constr ef in
let env' = Environ.push_named (LocalDef (f, ef, f_type)) env in
let evdref = ref sigma in
let suchthat = Constrintern.interp_type_evars env' evdref suchthat in
diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml
index 2b19c2805..92ece7ccf 100644
--- a/plugins/extraction/extraction.ml
+++ b/plugins/extraction/extraction.ml
@@ -42,11 +42,11 @@ let none = Evd.empty
let type_of env c =
let polyprop = (lang() == Haskell) in
- Retyping.get_type_of ~polyprop env none (strip_outer_cast c)
+ EConstr.Unsafe.to_constr (Retyping.get_type_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c)))
let sort_of env c =
let polyprop = (lang() == Haskell) in
- Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast c)
+ Retyping.get_sort_family_of ~polyprop env none (strip_outer_cast none (EConstr.of_constr c))
(*S Generation of flags and signatures. *)
@@ -70,11 +70,17 @@ type scheme = TypeScheme | Default
type flag = info * scheme
+let whd_all env t =
+ EConstr.Unsafe.to_constr (whd_all env none (EConstr.of_constr t))
+
+let whd_betaiotazeta t =
+ EConstr.Unsafe.to_constr (whd_betaiotazeta none (EConstr.of_constr t))
+
(*s [flag_of_type] transforms a type [t] into a [flag].
Really important function. *)
let rec flag_of_type env t : flag =
- let t = whd_all env none t in
+ let t = whd_all env t in
match kind_of_term 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)
@@ -99,17 +105,20 @@ let is_info_scheme env t = match flag_of_type env t with
| (Info, TypeScheme) -> true
| _ -> false
+let push_rel_assum (n, t) env =
+ Environ.push_rel (LocalAssum (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 none c) with
+ match kind_of_term (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 none c) with
+ match kind_of_term (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
@@ -135,7 +144,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 none c) with
+ match kind_of_term (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
@@ -143,7 +152,7 @@ let rec type_sign_vl env c =
| _ -> [],[]
let rec nb_default_params env c =
- match kind_of_term (whd_all env none c) with
+ match kind_of_term (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
@@ -214,7 +223,7 @@ let parse_ind_args si args relmax =
let rec extract_type env db j c args =
- match kind_of_term (whd_betaiotazeta Evd.empty c) with
+ match kind_of_term (whd_betaiotazeta c) with
| App (d, args') ->
(* We just accumulate the arguments. *)
extract_type env db j d (Array.to_list args' @ args)
@@ -297,7 +306,7 @@ and extract_type_app env db (r,s) args =
let ml_args =
List.fold_right
(fun (b,c) a -> if b == Keep then
- let p = List.length (fst (splay_prod env none (type_of env c))) in
+ let p = List.length (fst (splay_prod env none (EConstr.of_constr (type_of env c)))) in
let db = iterate (fun l -> 0 :: l) p db in
(extract_type_scheme env db c p) :: a
else a)
@@ -316,12 +325,13 @@ and extract_type_app env db (r,s) args =
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 Evd.empty c in
+ let c = whd_betaiotazeta c in
match kind_of_term c with
| Lambda (n,t,d) ->
extract_type_scheme (push_rel_assum (n,t) env) db d (p-1)
| _ ->
- let rels = fst (splay_prod env none (type_of env c)) in
+ let rels = fst (splay_prod env none (EConstr.of_constr (type_of env c))) in
+ let rels = List.map (on_snd EConstr.Unsafe.to_constr) rels in
let env = push_rels_assum rels env in
let eta_args = List.rev_map mkRel (List.interval 1 p) in
extract_type env db 0 (lift p c) eta_args
@@ -488,7 +498,7 @@ and extract_really_ind env kn mib =
*)
and extract_type_cons env db dbmap c i =
- match kind_of_term (whd_all env none c) with
+ match kind_of_term (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
@@ -595,7 +605,8 @@ let rec extract_term env mle mlt c args =
| Construct (cp,_) ->
extract_cons_app env mle mlt cp args
| Proj (p, c) ->
- let term = Retyping.expand_projection env (Evd.from_env env) p c [] in
+ let term = Retyping.expand_projection env (Evd.from_env env) p (EConstr.of_constr c) [] in
+ let term = EConstr.Unsafe.to_constr term in
extract_term env mle mlt term args
| Rel n ->
(* As soon as the expected [mlt] for the head is known, *)
@@ -846,8 +857,8 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt =
and decompose the term [c] in [n] lambdas, with eta-expansion if needed. *)
let decomp_lams_eta_n n m env c t =
- let rels = fst (splay_prod_n env none n t) in
- let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,c)) rels in
+ let rels = fst (splay_prod_n env none n (EConstr.of_constr t)) in
+ let rels = List.map (fun (LocalAssum (id,c) | LocalDef (id,_,c)) -> (id,EConstr.Unsafe.to_constr c)) rels in
let rels',c = decompose_lam c in
let d = n - m in
(* we'd better keep rels' as long as possible. *)
@@ -887,7 +898,7 @@ let extract_std_constant env kn body typ =
break user's clever let-ins and partial applications). *)
let rels, c =
let n = List.length s
- and m = nb_lam body in
+ and m = nb_lam Evd.empty (EConstr.of_constr body) (** FIXME *) in
if n <= m then decompose_lam_n n body
else
let s,s' = List.chop m s in
diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml
index b34a36492..7773f6a2f 100644
--- a/plugins/firstorder/formula.ml
+++ b/plugins/firstorder/formula.ml
@@ -54,7 +54,7 @@ let construct_nhyps ind gls =
let ind_hyps nevar ind largs gls=
let types= Inductiveops.arities_of_constructors (pf_env gls) ind in
let myhyps t =
- let t1=prod_applist t largs in
+ let t1=Term.prod_applist t largs in
let t2=snd (decompose_prod_n_assum nevar t1) in
fst (decompose_prod_assum t2) in
Array.map myhyps types
@@ -76,18 +76,22 @@ type kind_of_formula=
| Forall of constr*constr
| Atom of constr
+let pop t = Vars.lift (-1) t
+
let kind_of_formula gl term =
let normalize=special_nf gl in
let cciterm=special_whd gl term in
- match match_with_imp_term cciterm with
- Some (a,b)-> Arrow(a,(pop b))
+ match match_with_imp_term (project gl) (EConstr.of_constr cciterm) with
+ Some (a,b)-> Arrow(EConstr.Unsafe.to_constr a,(pop (EConstr.Unsafe.to_constr b)))
|_->
- match match_with_forall_term cciterm with
- Some (_,a,b)-> Forall(a,b)
+ match match_with_forall_term (project gl) (EConstr.of_constr cciterm) with
+ Some (_,a,b)-> Forall(EConstr.Unsafe.to_constr a,EConstr.Unsafe.to_constr b)
|_->
- match match_with_nodep_ind cciterm with
+ match match_with_nodep_ind (project gl) (EConstr.of_constr cciterm) with
Some (i,l,n)->
- let ind,u=destInd i in
+ let l = List.map EConstr.Unsafe.to_constr l in
+ let ind,u=EConstr.destInd (project gl) i in
+ let u = EConstr.EInstance.kind (project gl) u in
let (mib,mip) = Global.lookup_inductive ind in
let nconstr=Array.length mip.mind_consnames in
if Int.equal nconstr 0 then
@@ -96,7 +100,7 @@ let kind_of_formula gl term =
let has_realargs=(n>0) in
let is_trivial=
let is_constant c =
- Int.equal (nb_prod c) mib.mind_nparams in
+ Int.equal (nb_prod (project gl) (EConstr.of_constr c)) mib.mind_nparams in
Array.exists is_constant mip.mind_nf_lc in
if Inductiveops.mis_is_recursive (ind,mib,mip) ||
(has_realargs && not is_trivial)
@@ -108,8 +112,11 @@ let kind_of_formula gl term =
else
Or((ind,u),l,is_trivial)
| _ ->
- match match_with_sigma_type cciterm with
- Some (i,l)-> Exists((destInd i),l)
+ match match_with_sigma_type (project gl) (EConstr.of_constr cciterm) with
+ Some (i,l)->
+ let (ind, u) = EConstr.destInd (project gl) i in
+ let u = EConstr.EInstance.kind (project gl) u in
+ Exists((ind, u), List.map EConstr.Unsafe.to_constr l)
|_-> Atom (normalize cciterm)
type atoms = {positive:constr list;negative:constr list}
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index eebd974ea..9dc2a51a6 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -107,7 +107,7 @@ let mk_open_instance id idc gl m t=
let typ=pf_unsafe_type_of gl idc in
(* since we know we will get a product,
reduction is not too expensive *)
- let (nam,_,_)=destProd (whd_all env evmap typ) in
+ let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap typ)) in
match nam with
Name id -> id
| Anonymous -> dummy_bvid in
@@ -119,19 +119,20 @@ let mk_open_instance id idc gl m t=
let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in
let evmap = Sigma.to_evar_map evmap in
let decl = LocalAssum (Name nid, c) in
- aux (n-1) (nid::avoid) (Environ.push_rel decl env) evmap (decl::decls) in
+ aux (n-1) (nid::avoid) (EConstr.push_rel decl env) evmap (decl::decls) in
let evmap, decls = aux m [] env evmap [] in
evmap, decls, revt
(* tactics *)
let left_instance_tac (inst,id) continue seq=
+ let open EConstr in
match inst with
Phantom dom->
if lookup (id,None) seq then
tclFAIL 0 (Pp.str "already done")
else
- tclTHENS (Proofview.V82.of_tactic (cut dom))
+ tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom)))
[tclTHENLIST
[Proofview.V82.of_tactic introf;
pf_constr_of_global id (fun idc ->
@@ -151,6 +152,7 @@ let left_instance_tac (inst,id) continue seq=
pf_constr_of_global id (fun idc ->
fun gl->
let evmap,rc,ot = mk_open_instance id idc gl m t in
+ let ot = EConstr.of_constr ot in
let gt=
it_mkLambda_or_LetIn
(mkApp(idc,[|ot|])) rc in
@@ -160,6 +162,7 @@ let left_instance_tac (inst,id) continue seq=
error "Untypable instance, maybe higher-order non-prenex quantification" in
tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl)
else
+ let t = EConstr.of_constr t in
pf_constr_of_global id (fun idc ->
Proofview.V82.of_tactic (generalize [mkApp(idc,[|t|])]))
in
@@ -172,16 +175,16 @@ let left_instance_tac (inst,id) continue seq=
let right_instance_tac inst continue seq=
match inst with
Phantom dom ->
- tclTHENS (Proofview.V82.of_tactic (cut dom))
+ tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr dom)))
[tclTHENLIST
[Proofview.V82.of_tactic introf;
(fun gls->
Proofview.V82.of_tactic (split (ImplicitBindings
- [mkVar (Tacmach.pf_nth_hyp_id gls 1)])) gls);
+ [EConstr.mkVar (Tacmach.pf_nth_hyp_id gls 1)])) gls);
tclSOLVE [wrap 0 true continue (deepen seq)]];
tclTRY (Proofview.V82.of_tactic assumption)]
| Real ((0,t),_) ->
- (tclTHEN (Proofview.V82.of_tactic (split (ImplicitBindings [t])))
+ (tclTHEN (Proofview.V82.of_tactic (split (ImplicitBindings [EConstr.of_constr t])))
(tclSOLVE [wrap 0 true continue (deepen seq)]))
| Real ((m,t),_) ->
tclFAIL 0 (Pp.str "not implemented ... yet")
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index 7ffc78928..a60fd4d8f 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -38,14 +38,14 @@ let wrap n b continue seq gls=
[]->anomaly (Pp.str "Not the expected number of hyps")
| nd::q->
let id = NamedDecl.get_id nd in
- if occur_var env id (pf_concl gls) ||
- List.exists (occur_var_in_decl env id) ctx then
+ if occur_var env (project gls) id (pf_concl gls) ||
+ List.exists (occur_var_in_decl env (project gls) id) ctx then
(aux (i-1) q (nd::ctx))
else
- add_formula Hyp (VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) gls in
+ add_formula Hyp (VarRef id) (EConstr.Unsafe.to_constr (NamedDecl.get_type nd)) (aux (i-1) q (nd::ctx)) gls in
let seq1=aux n nc [] in
let seq2=if b then
- add_formula Concl dummy_id (pf_concl gls) seq1 gls else seq1 in
+ add_formula Concl dummy_id (EConstr.Unsafe.to_constr (pf_concl gls)) seq1 gls else seq1 in
continue seq2 gls
let basename_of_global=function
@@ -63,12 +63,13 @@ let axiom_tac t seq=
with Not_found->tclFAIL 0 (Pp.str "No axiom link")
let ll_atom_tac a backtrack id continue seq=
+ let open EConstr in
tclIFTHENELSE
(try
tclTHENLIST
[pf_constr_of_global (find_left a seq) (fun left ->
pf_constr_of_global id (fun id ->
- Proofview.V82.of_tactic (generalize [mkApp(id, [|left|])])));
+ Proofview.V82.of_tactic (generalize [(mkApp(id, [|left|]))])));
clear_global id;
Proofview.V82.of_tactic intro]
with Not_found->tclFAIL 0 (Pp.str "No link"))
@@ -131,9 +132,9 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl=
let vars=Array.init p (fun j->mkRel (p-j)) in
let capply=mkApp ((lift p cstr),vars) in
let head=mkApp ((lift p idc),[|capply|]) in
- it_mkLambda_or_LetIn head rc in
+ EConstr.of_constr (it_mkLambda_or_LetIn head rc) in
let lp=Array.length rcs in
- let newhyps idc =List.init lp (myterm idc) in
+ let newhyps idc =List.init lp (myterm (EConstr.Unsafe.to_constr idc)) in
tclIFTHENELSE
(tclTHENLIST
[pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize (newhyps idc)));
@@ -142,8 +143,13 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl=
(wrap lp false continue seq) backtrack gl
let ll_arrow_tac a b c backtrack id continue seq=
+ let open EConstr in
+ let open Vars in
+ let a = EConstr.of_constr a in
+ let b = EConstr.of_constr b in
+ let c = EConstr.of_constr c in
let cc=mkProd(Anonymous,a,(lift 1 b)) in
- let d idc =mkLambda (Anonymous,b,
+ let d idc = mkLambda (Anonymous,b,
mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in
tclORELSE
(tclTHENS (Proofview.V82.of_tactic (cut c))
@@ -186,11 +192,12 @@ let left_exists_tac ind backtrack id continue seq gls=
let ll_forall_tac prod backtrack id continue seq=
tclORELSE
- (tclTHENS (Proofview.V82.of_tactic (cut prod))
+ (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr prod)))
[tclTHENLIST
[Proofview.V82.of_tactic intro;
pf_constr_of_global id (fun idc ->
(fun gls->
+ let open EConstr in
let id0=pf_nth_hyp_id gls 1 in
let term=mkApp(idc,[|mkVar(id0)|]) in
tclTHEN (Proofview.V82.of_tactic (generalize [term])) (Proofview.V82.of_tactic (clear [id0])) gls));
diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml
index 1248b60a7..fb0c22c2b 100644
--- a/plugins/firstorder/sequent.ml
+++ b/plugins/firstorder/sequent.ml
@@ -200,7 +200,8 @@ let extend_with_ref_list l seq gl =
let l = expand_constructor_hints l in
let f gr (seq,gl) =
let gl, c = pf_eapply Evd.fresh_global gl gr in
- let typ=(pf_unsafe_type_of gl c) in
+ let typ=(pf_unsafe_type_of gl (EConstr.of_constr c)) in
+ let typ = EConstr.Unsafe.to_constr typ in
(add_formula Hyp gr typ seq gl,gl) in
List.fold_right f l (seq,gl)
@@ -214,8 +215,9 @@ let extend_with_auto_hints l seq gl=
| Res_pf_THEN_trivial_fail (c,_) ->
let (c, _, _) = c in
(try
- let gr = global_of_constr c in
+ let (gr, _) = Termops.global_of_constr (project gl) c in
let typ=(pf_unsafe_type_of gl c) in
+ let typ = EConstr.Unsafe.to_constr typ in
seqref:=add_formula Hint gr typ !seqref gl
with Not_found->())
| _-> () in
diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml
index d9ab36ad6..7cbfb8e7d 100644
--- a/plugins/firstorder/unify.ml
+++ b/plugins/firstorder/unify.ml
@@ -21,7 +21,13 @@ exception UFAIL of constr*constr
to the equation set. Raises UFAIL with a pair of terms
*)
+let strip_outer_cast t =
+ EConstr.Unsafe.to_constr (strip_outer_cast Evd.empty (EConstr.of_constr t)) (** FIXME *)
+
+let pop t = Vars.lift (-1) t
+
let unif t1 t2=
+ let evd = Evd.empty in (** FIXME *)
let bige=Queue.create ()
and sigma=ref [] in
let bind i t=
@@ -38,8 +44,8 @@ let unif t1 t2=
Queue.add (t1,t2) bige;
try while true do
let t1,t2=Queue.take bige in
- let nt1=head_reduce (whd_betaiotazeta Evd.empty t1)
- and nt2=head_reduce (whd_betaiotazeta Evd.empty t2) in
+ let nt1=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t1)))
+ and nt2=head_reduce (EConstr.Unsafe.to_constr (whd_betaiotazeta evd (EConstr.of_constr t2))) in
match (kind_of_term nt1),(kind_of_term nt2) with
Meta i,Meta j->
if not (Int.equal i j) then
@@ -47,13 +53,13 @@ let unif t1 t2=
else bind i nt2
| Meta i,_ ->
let t=subst_meta !sigma nt2 in
- if Int.Set.is_empty (free_rels t) &&
- not (occur_term (mkMeta i) t) then
+ if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) &&
+ not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then
bind i t else raise (UFAIL(nt1,nt2))
| _,Meta i ->
let t=subst_meta !sigma nt1 in
- if Int.Set.is_empty (free_rels t) &&
- not (occur_term (mkMeta i) t) then
+ if Int.Set.is_empty (free_rels evd (EConstr.of_constr t)) &&
+ not (occur_term evd (EConstr.mkMeta i) (EConstr.of_constr t)) then
bind i t else raise (UFAIL(nt1,nt2))
| Cast(_,_,_),_->Queue.add (strip_outer_cast nt1,nt2) bige
| _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast nt2) bige
diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml
index 8e193c753..e11cbc279 100644
--- a/plugins/fourier/fourierR.ml
+++ b/plugins/fourier/fourierR.ml
@@ -190,6 +190,8 @@ type hineq={hname:constr; (* le nom de l'hypothèse *)
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
| App (f,args) ->
(match kind_of_term f with
@@ -281,6 +283,8 @@ let fourier_lineq lineq1 =
(* Defined constants *)
let get = Lazy.force
+let cget = get
+let eget c = EConstr.of_constr (Lazy.force c)
let constant = Coqlib.gen_constant "Fourier"
(* Standard library *)
@@ -373,6 +377,7 @@ let rational_to_real x =
(* preuve que 0<n*1/d
*)
let tac_zero_inf_pos gl (n,d) =
+ let get = eget in
let tacn=ref (apply (get coq_Rlt_zero_1)) in
let tacd=ref (apply (get coq_Rlt_zero_1)) in
for _i = 1 to n - 1 do
@@ -385,6 +390,7 @@ let tac_zero_inf_pos gl (n,d) =
(* preuve que 0<=n*1/d
*)
let tac_zero_infeq_pos gl (n,d)=
+ let get = eget in
let tacn=ref (if n=0
then (apply (get coq_Rle_zero_zero))
else (apply (get coq_Rle_zero_1))) in
@@ -399,7 +405,8 @@ let tac_zero_infeq_pos gl (n,d)=
(* preuve que 0<(-n)*(1/d) => False
*)
let tac_zero_inf_false gl (n,d) =
- if n=0 then (apply (get coq_Rnot_lt0))
+ let get = eget in
+if n=0 then (apply (get coq_Rnot_lt0))
else
(Tacticals.New.tclTHEN (apply (get coq_Rle_not_lt))
(tac_zero_infeq_pos gl (-n,d)))
@@ -408,6 +415,7 @@ let tac_zero_inf_false gl (n,d) =
(* preuve que 0<=(-n)*(1/d) => False
*)
let tac_zero_infeq_false gl (n,d) =
+ let get = eget in
(Tacticals.New.tclTHEN (apply (get coq_Rlt_not_le_frac_opp))
(tac_zero_inf_pos gl (-n,d)))
;;
@@ -415,7 +423,8 @@ let tac_zero_infeq_false gl (n,d) =
let exact = exact_check;;
let tac_use h =
- let tac = exact h.hname in
+ let get = eget in
+ let tac = exact (EConstr.of_constr h.hname) in
match h.htype with
"Rlt" -> tac
|"Rle" -> tac
@@ -461,14 +470,17 @@ exception GoalDone
let rec fourier () =
Proofview.Goal.nf_enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
Coqlib.check_required_library ["Coq";"fourier";"Fourier"];
- let goal = Termops.strip_outer_cast concl in
+ let goal = Termops.strip_outer_cast sigma concl in
+ let goal = EConstr.Unsafe.to_constr goal in
let fhyp=Id.of_string "new_hyp_for_fourier" in
(* si le but est une inéquation, on introduit son contraire,
et le but à prouver devient False *)
try
match (kind_of_term goal) with
App (f,args) ->
+ let get = eget in
(match (string_of_R_constr f) with
"Rlt" ->
(Tacticals.New.tclTHEN
@@ -494,7 +506,7 @@ let rec fourier () =
|_-> raise GoalDone
with GoalDone ->
(* les hypothèses *)
- let hyps = List.map (fun (h,t)-> (mkVar h,t))
+ let hyps = List.map (fun (h,t)-> (EConstr.mkVar h,t))
(list_of_sign (Proofview.Goal.hyps gl)) in
let lineq =ref [] in
List.iter (fun h -> try (lineq:=(ineq1_of_constr h)@(!lineq))
@@ -547,6 +559,7 @@ let rec fourier () =
!t2 |] in
let tc=rational_to_real cres in
(* puis sa preuve *)
+ let get = eget in
let tac1=ref (if h1.hstrict
then (Tacticals.New.tclTHENS (apply (get coq_Rfourier_lt))
[tac_use h1;
@@ -583,28 +596,28 @@ let rec fourier () =
then tac_zero_inf_false gl (rational_to_fraction cres)
else tac_zero_infeq_false gl (rational_to_fraction cres)
in
- tac:=(Tacticals.New.tclTHENS (cut ineq)
+ tac:=(Tacticals.New.tclTHENS (cut (EConstr.of_constr ineq))
[Tacticals.New.tclTHEN (change_concl
- (mkAppL [| get coq_not; ineq|]
- ))
+ (EConstr.of_constr (mkAppL [| cget coq_not; ineq|]
+ )))
(Tacticals.New.tclTHEN (apply (if sres then get coq_Rnot_lt_lt
else get coq_Rnot_le_le))
(Tacticals.New.tclTHENS (Equality.replace
- (mkAppL [|get coq_Rminus;!t2;!t1|]
- )
- tc)
+ (EConstr.of_constr (mkAppL [|cget coq_Rminus;!t2;!t1|]
+ ))
+ (EConstr.of_constr tc))
[tac2;
(Tacticals.New.tclTHENS
(Equality.replace
- (mkApp (get coq_Rinv,
- [|get coq_R1|]))
+ (EConstr.of_constr (mkApp (cget coq_Rinv,
+ [|cget coq_R1|])))
(get coq_R1))
(* en attendant Field, ça peut aider Ring de remplacer 1/1 par 1 ... *)
[Tacticals.New.tclORELSE
(* TODO : Ring.polynom []*) (Proofview.tclUNIT ())
(Proofview.tclUNIT ());
- Tacticals.New.pf_constr_of_global (get coq_sym_eqT) (fun symeq ->
+ Tacticals.New.pf_constr_of_global (cget coq_sym_eqT) (fun symeq ->
(Tacticals.New.tclTHEN (apply symeq)
(apply (get coq_Rinv_1))))]
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 3199474dd..48c0f5f04 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -2,6 +2,7 @@ open Printer
open CErrors
open Util
open Term
+open EConstr
open Vars
open Namegen
open Names
@@ -18,6 +19,12 @@ open Context.Rel.Declaration
module RelDecl = Context.Rel.Declaration
+let local_assum (na, t) =
+ RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t)
+
+let local_def (na, b, t) =
+ RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t)
+
(* let msgnl = Pp.msgnl *)
(*
@@ -95,6 +102,7 @@ let list_chop ?(msg="") n l =
with Failure (msg') ->
failwith (msg ^ msg')
+let pop t = Vars.lift (-1) t
let make_refl_eq constructor type_of_t t =
(* let refl_equal_term = Lazy.force refl_equal in *)
@@ -131,16 +139,16 @@ let refine c =
let thin l = Proofview.V82.of_tactic (Tactics.clear l)
-let eq_constr u v = eq_constr_nounivs u v
+let eq_constr sigma u v = EConstr.eq_constr_nounivs sigma u v
-let is_trivial_eq t =
+let is_trivial_eq sigma t =
let res = try
begin
- match kind_of_term t with
- | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) ->
- eq_constr t1 t2
- | App(f,[|t1;a1;t2;a2|]) when eq_constr f (jmeq ()) ->
- eq_constr t1 t2 && eq_constr a1 a2
+ match EConstr.kind sigma t with
+ | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) ->
+ eq_constr sigma t1 t2
+ | App(f,[|t1;a1;t2;a2|]) when eq_constr sigma f (jmeq ()) ->
+ eq_constr sigma t1 t2 && eq_constr sigma a1 a2
| _ -> false
end
with e when CErrors.noncritical e -> false
@@ -148,30 +156,30 @@ let is_trivial_eq t =
(* observe (str "is_trivial_eq " ++ Printer.pr_lconstr t ++ (if res then str " true" else str " false")); *)
res
-let rec incompatible_constructor_terms t1 t2 =
- let c1,arg1 = decompose_app t1
- and c2,arg2 = decompose_app t2
+let rec incompatible_constructor_terms sigma t1 t2 =
+ let c1,arg1 = decompose_app sigma t1
+ and c2,arg2 = decompose_app sigma t2
in
- (not (eq_constr t1 t2)) &&
- isConstruct c1 && isConstruct c2 &&
+ (not (eq_constr sigma t1 t2)) &&
+ isConstruct sigma c1 && isConstruct sigma c2 &&
(
- not (eq_constr c1 c2) ||
- List.exists2 incompatible_constructor_terms arg1 arg2
+ not (eq_constr sigma c1 c2) ||
+ List.exists2 (incompatible_constructor_terms sigma) arg1 arg2
)
-let is_incompatible_eq t =
+let is_incompatible_eq sigma t =
let res =
try
- match kind_of_term t with
- | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) ->
- incompatible_constructor_terms t1 t2
- | App(f,[|u1;t1;u2;t2|]) when eq_constr f (jmeq ()) ->
- (eq_constr u1 u2 &&
- incompatible_constructor_terms t1 t2)
+ match EConstr.kind sigma t with
+ | App(f,[|_;t1;t2|]) when eq_constr sigma f (Lazy.force eq) ->
+ incompatible_constructor_terms sigma t1 t2
+ | App(f,[|u1;t1;u2;t2|]) when eq_constr sigma f (jmeq ()) ->
+ (eq_constr sigma u1 u2 &&
+ incompatible_constructor_terms sigma t1 t2)
| _ -> false
with e when CErrors.noncritical e -> false
in
- if res then observe (str "is_incompatible_eq " ++ Printer.pr_lconstr t);
+ if res then observe (str "is_incompatible_eq " ++ Printer.pr_leconstr t);
res
let change_hyp_with_using msg hyp_id t tac : tactic =
@@ -208,40 +216,38 @@ let prove_trivial_eq h_id context (constructor,type_of_term,term) =
-let find_rectype env c =
- let (t, l) = decompose_app (Reduction.whd_betaiotazeta env c) in
- match kind_of_term t with
+let find_rectype env sigma c =
+ let (t, l) = decompose_app sigma (Reductionops.whd_betaiotazeta sigma c) in
+ match EConstr.kind sigma t with
| Ind ind -> (t, l)
| Construct _ -> (t,l)
| _ -> raise Not_found
-let isAppConstruct ?(env=Global.env ()) t =
+let isAppConstruct ?(env=Global.env ()) sigma t =
try
- let t',l = find_rectype (Global.env ()) t in
- observe (str "isAppConstruct : " ++ Printer.pr_lconstr t ++ str " -> " ++ Printer.pr_lconstr (applist (t',l)));
+ let t',l = find_rectype env sigma t in
+ observe (str "isAppConstruct : " ++ Printer.pr_leconstr t ++ str " -> " ++ Printer.pr_leconstr (applist (t',l)));
true
with Not_found -> false
let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *)
- let clos_norm_flags flgs env sigma t =
- CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
+ Reductionops.clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
-let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
+let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
let nochange ?t' msg =
begin
- observe (str ("Not treating ( "^msg^" )") ++ pr_lconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_lconstr t );
+ observe (str ("Not treating ( "^msg^" )") ++ pr_leconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr t );
failwith "NoChange";
end
in
- let eq_constr = Evarconv.e_conv env (ref sigma) in
- if not (noccurn 1 end_of_type)
+ let eq_constr c1 c2 = Evarconv.e_conv env (ref sigma) c1 c2 in
+ if not (noccurn sigma 1 end_of_type)
then nochange "dependent"; (* if end_of_type depends on this term we don't touch it *)
- if not (isApp t) then nochange "not an equality";
- let f_eq,args = destApp t in
+ if not (isApp sigma t) then nochange "not an equality";
+ let f_eq,args = destApp sigma t in
let constructor,t1,t2,t1_typ =
try
if (eq_constr f_eq (Lazy.force eq))
@@ -258,36 +264,36 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
else nochange "not an equality"
with e when CErrors.noncritical e -> nochange "not an equality"
in
- if not ((closed0 (fst t1)) && (closed0 (snd t1)))then nochange "not a closed lhs";
+ if not ((closed0 sigma (fst t1)) && (closed0 sigma (snd t1)))then nochange "not a closed lhs";
let rec compute_substitution sub t1 t2 =
(* observe (str "compute_substitution : " ++ pr_lconstr t1 ++ str " === " ++ pr_lconstr t2); *)
- if isRel t2
+ if isRel sigma t2
then
- let t2 = destRel t2 in
+ let t2 = destRel sigma t2 in
begin
try
let t1' = Int.Map.find t2 sub in
if not (eq_constr t1 t1') then nochange "twice bound variable";
sub
with Not_found ->
- assert (closed0 t1);
+ assert (closed0 sigma t1);
Int.Map.add t2 t1 sub
end
- else if isAppConstruct t1 && isAppConstruct t2
+ else if isAppConstruct sigma t1 && isAppConstruct sigma t2
then
begin
- let c1,args1 = find_rectype env t1
- and c2,args2 = find_rectype env t2
+ let c1,args1 = find_rectype env sigma t1
+ and c2,args2 = find_rectype env sigma t2
in
if not (eq_constr c1 c2) then nochange "cannot solve (diff)";
List.fold_left2 compute_substitution sub args1 args2
end
else
- if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reduction.whd_all env t1) t2) "cannot solve (diff)"
+ if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reductionops.whd_all env sigma t1) t2) "cannot solve (diff)"
in
let sub = compute_substitution Int.Map.empty (snd t1) (snd t2) in
let sub = compute_substitution sub (fst t1) (fst t2) in
- let end_of_type_with_pop = Termops.pop end_of_type in (*the equation will be removed *)
+ let end_of_type_with_pop = pop end_of_type in (*the equation will be removed *)
let new_end_of_type =
(* Ugly hack to prevent Map.fold order change between ocaml-3.08.3 and ocaml-3.08.4
Can be safely replaced by the next comment for Ocaml >= 3.08.4
@@ -309,7 +315,7 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
try
let witness = Int.Map.find i sub in
if is_local_def decl then anomaly (Pp.str "can not redefine a rel!");
- (Termops.pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun))
+ (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, RelDecl.get_type decl, witness_fun))
with Not_found ->
(mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun)
)
@@ -318,9 +324,9 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
context
in
let new_type_of_hyp =
- Reductionops.nf_betaiota Evd.empty new_type_of_hyp in
+ Reductionops.nf_betaiota sigma new_type_of_hyp in
let new_ctxt,new_end_of_type =
- decompose_prod_n_assum ctxt_size new_type_of_hyp
+ decompose_prod_n_assum sigma ctxt_size new_type_of_hyp
in
let prove_new_hyp : tactic =
tclTHEN
@@ -353,21 +359,21 @@ let change_eq env sigma hyp_id (context:Context.Rel.t) x t end_of_type =
new_ctxt,new_end_of_type,simpl_eq_tac
-let is_property (ptes_info:ptes_info) t_x full_type_of_hyp =
- if isApp t_x
+let is_property sigma (ptes_info:ptes_info) t_x full_type_of_hyp =
+ if isApp sigma t_x
then
- let pte,args = destApp t_x in
- if isVar pte && Array.for_all closed0 args
+ let pte,args = destApp sigma t_x in
+ if isVar sigma pte && Array.for_all (closed0 sigma) args
then
try
- let info = Id.Map.find (destVar pte) ptes_info in
+ let info = Id.Map.find (destVar sigma pte) ptes_info in
info.is_valid full_type_of_hyp
with Not_found -> false
else false
else false
-let isLetIn t =
- match kind_of_term t with
+let isLetIn sigma t =
+ match EConstr.kind sigma t with
| LetIn _ -> true
| _ -> false
@@ -387,8 +393,9 @@ let rewrite_until_var arg_num eq_ids : tactic =
will break the Guard when trying to save the Lemma.
*)
let test_var g =
- let _,args = destApp (pf_concl g) in
- not ((isConstruct args.(arg_num)) || isAppConstruct args.(arg_num))
+ let sigma = project g in
+ let _,args = destApp sigma (pf_concl g) in
+ not ((isConstruct sigma args.(arg_num)) || isAppConstruct sigma args.(arg_num))
in
let rec do_rewrite eq_ids g =
if test_var g
@@ -407,30 +414,30 @@ let rewrite_until_var arg_num eq_ids : tactic =
let rec_pte_id = Id.of_string "Hrec"
let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
- let coq_False = Coqlib.build_coq_False () in
- let coq_True = Coqlib.build_coq_True () in
- let coq_I = Coqlib.build_coq_I () in
+ let coq_False = EConstr.of_constr (Coqlib.build_coq_False ()) in
+ let coq_True = EConstr.of_constr (Coqlib.build_coq_True ()) in
+ let coq_I = EConstr.of_constr (Coqlib.build_coq_I ()) in
let rec scan_type context type_of_hyp : tactic =
- if isLetIn type_of_hyp then
+ if isLetIn sigma type_of_hyp then
let real_type_of_hyp = it_mkProd_or_LetIn type_of_hyp context in
let reduced_type_of_hyp = nf_betaiotazeta real_type_of_hyp in
(* length of context didn't change ? *)
let new_context,new_typ_of_hyp =
- decompose_prod_n_assum (List.length context) reduced_type_of_hyp
+ decompose_prod_n_assum sigma (List.length context) reduced_type_of_hyp
in
tclTHENLIST
[ h_reduce_with_zeta (Locusops.onHyp hyp_id);
scan_type new_context new_typ_of_hyp ]
- else if isProd type_of_hyp
+ else if isProd sigma type_of_hyp
then
begin
- let (x,t_x,t') = destProd type_of_hyp in
+ let (x,t_x,t') = destProd sigma type_of_hyp in
let actual_real_type_of_hyp = it_mkProd_or_LetIn t' context in
- if is_property ptes_infos t_x actual_real_type_of_hyp then
+ if is_property sigma ptes_infos t_x actual_real_type_of_hyp then
begin
- let pte,pte_args = (destApp t_x) in
- let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar pte) ptes_infos).proving_tac in
- let popped_t' = Termops.pop t' in
+ let pte,pte_args = (destApp sigma t_x) in
+ let (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar sigma pte) ptes_infos).proving_tac in
+ let popped_t' = pop t' in
let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in
let prove_new_type_of_hyp =
let context_length = List.length context in
@@ -467,20 +474,20 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
scan_type context popped_t'
]
end
- else if eq_constr t_x coq_False then
+ else if eq_constr sigma t_x coq_False then
begin
(* observe (str "Removing : "++ Ppconstr.pr_id hyp_id++ *)
(* str " since it has False in its preconds " *)
(* ); *)
raise TOREMOVE; (* False -> .. useless *)
end
- else if is_incompatible_eq t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *)
- else if eq_constr t_x coq_True (* Trivial => we remove this precons *)
+ else if is_incompatible_eq sigma t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *)
+ else if eq_constr sigma t_x coq_True (* Trivial => we remove this precons *)
then
(* observe (str "In "++Ppconstr.pr_id hyp_id++ *)
(* str " removing useless precond True" *)
(* ); *)
- let popped_t' = Termops.pop t' in
+ let popped_t' = pop t' in
let real_type_of_hyp =
it_mkProd_or_LetIn popped_t' context
in
@@ -506,15 +513,15 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
((* observe_tac "prove_trivial" *) prove_trivial);
scan_type context popped_t'
]
- else if is_trivial_eq t_x
+ else if is_trivial_eq sigma t_x
then (* t_x := t = t => we remove this precond *)
- let popped_t' = Termops.pop t' in
+ let popped_t' = pop t' in
let real_type_of_hyp =
it_mkProd_or_LetIn popped_t' context
in
- let hd,args = destApp t_x in
+ let hd,args = destApp sigma t_x in
let get_args hd args =
- if eq_constr hd (Lazy.force eq)
+ if eq_constr sigma hd (Lazy.force eq)
then (Lazy.force refl_equal,args.(0),args.(1))
else (jmeq_refl (),args.(0),args.(1))
in
@@ -597,18 +604,18 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
let new_term_value_eq = pf_unsafe_type_of g' (mkVar heq_id) in
(* compute the new value of the body *)
let new_term_value =
- match kind_of_term new_term_value_eq with
+ match EConstr.kind (project g') new_term_value_eq with
| App(f,[| _;_;args2 |]) -> args2
| _ ->
observe (str "cannot compute new term value : " ++ pr_gls g' ++ fnl () ++ str "last hyp is" ++
- pr_lconstr_env (pf_env g') Evd.empty new_term_value_eq
+ pr_leconstr_env (pf_env g') (project g') new_term_value_eq
);
anomaly (Pp.str "cannot compute new term value")
in
let fun_body =
mkLambda(Anonymous,
pf_unsafe_type_of g' term,
- Termops.replace_term term (mkRel 1) dyn_infos.info
+ Termops.replace_term (project g') term (mkRel 1) dyn_infos.info
)
in
let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in
@@ -691,15 +698,16 @@ let build_proof
: tactic =
let rec build_proof_aux do_finalize dyn_infos : tactic =
fun g ->
+ let sigma = project g in
(* observe (str "proving on " ++ Printer.pr_lconstr_env (pf_env g) term);*)
- match kind_of_term dyn_infos.info with
+ match EConstr.kind sigma dyn_infos.info with
| Case(ci,ct,t,cb) ->
let do_finalize_t dyn_info' =
fun g ->
let t = dyn_info'.info in
let dyn_infos = {dyn_info' with info =
mkCase(ci,ct,t,cb)} in
- let g_nb_prod = nb_prod (pf_concl g) in
+ let g_nb_prod = nb_prod (project g) (pf_concl g) in
let type_of_term = pf_unsafe_type_of g t in
let term_eq =
make_refl_eq (Lazy.force refl_equal) type_of_term t
@@ -712,7 +720,7 @@ let build_proof
(fun g -> observe_tac "toto" (
tclTHENSEQ [Proofview.V82.of_tactic (Simple.case t);
(fun g' ->
- let g'_nb_prod = nb_prod (pf_concl g') in
+ let g'_nb_prod = nb_prod (project g') (pf_concl g') in
let nb_instanciate_partial = g'_nb_prod - g_nb_prod in
observe_tac "treat_new_case"
(treat_new_case
@@ -732,7 +740,7 @@ let build_proof
build_proof do_finalize_t {dyn_infos with info = t} g
| Lambda(n,t,b) ->
begin
- match kind_of_term( pf_concl g) with
+ match EConstr.kind sigma (pf_concl g) with
| Prod _ ->
tclTHEN
(Proofview.V82.of_tactic intro)
@@ -762,9 +770,9 @@ let build_proof
| Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ ->
do_finalize dyn_infos g
| App(_,_) ->
- let f,args = decompose_app dyn_infos.info in
+ let f,args = decompose_app sigma dyn_infos.info in
begin
- match kind_of_term f with
+ match EConstr.kind sigma f with
| App _ -> assert false (* we have collected all the app in decompose_app *)
| Proj _ -> assert false (*FIXME*)
| Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ ->
@@ -786,7 +794,7 @@ let build_proof
do_finalize dyn_infos g
| Lambda _ ->
let new_term =
- Reductionops.nf_beta Evd.empty dyn_infos.info in
+ Reductionops.nf_beta sigma dyn_infos.info in
build_proof do_finalize {dyn_infos with info = new_term}
g
| LetIn _ ->
@@ -838,7 +846,7 @@ let build_proof
| Rel _ -> anomaly (Pp.str "Free var in goal conclusion !")
and build_proof do_finalize dyn_infos g =
(* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *)
- observe_tac_stream (str "build_proof with " ++ Printer.pr_lconstr dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g
+ observe_tac_stream (str "build_proof with " ++ Printer.pr_leconstr dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g
and build_proof_args do_finalize dyn_infos (* f_args' args *) :tactic =
fun g ->
let (f_args',args) = dyn_infos.info in
@@ -904,7 +912,7 @@ let prove_rec_hyp_for_struct fix_info =
(fun eq_hyps -> tclTHEN
(rewrite_until_var (fix_info.idx) eq_hyps)
(fun g ->
- let _,pte_args = destApp (pf_concl g) in
+ let _,pte_args = destApp (project g) (pf_concl g) in
let rec_hyp_proof =
mkApp(mkVar fix_info.name,array_get_start pte_args)
in
@@ -925,10 +933,11 @@ let generalize_non_dep hyp g =
let to_revert,_ =
let open Context.Named.Declaration in
Environ.fold_named_context_reverse (fun (clear,keep) decl ->
+ let decl = map_named_decl EConstr.of_constr decl in
let hyp = get_id decl in
if Id.List.mem hyp hyps
- || List.exists (Termops.occur_var_in_decl env hyp) keep
- || Termops.occur_var env hyp hyp_typ
+ || List.exists (Termops.occur_var_in_decl env (project g) hyp) keep
+ || Termops.occur_var env (project g) hyp hyp_typ
|| Termops.is_section_variable hyp (* should be dangerous *)
then (clear,decl::keep)
else (hyp::clear,keep))
@@ -951,11 +960,12 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num
(* observe (str "nb_args := " ++ str (string_of_int nb_args)); *)
(* observe (str "nb_params := " ++ str (string_of_int nb_params)); *)
(* observe (str "rec_args_num := " ++ str (string_of_int (rec_args_num + 1) )); *)
- let f_def = Global.lookup_constant (fst (destConst f)) in
+ let f_def = Global.lookup_constant (fst (destConst evd f)) in
let eq_lhs = mkApp(f,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i))) in
let f_body = Option.get (Global.body_of_constant_body f_def) in
- let params,f_body_with_params = decompose_lam_n nb_params f_body in
- let (_,num),(_,_,bodies) = destFix f_body_with_params in
+ let f_body = EConstr.of_constr f_body in
+ let params,f_body_with_params = decompose_lam_n evd nb_params f_body in
+ let (_,num),(_,_,bodies) = destFix evd f_body_with_params in
let fnames_with_params =
let params = Array.init nb_params (fun i -> mkRel(nb_params - i)) in
let fnames = List.rev (Array.to_list (Array.map (fun f -> mkApp(f,params)) fnames)) in
@@ -970,13 +980,13 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num
let (type_ctxt,type_of_f),evd =
let evd,t = Typing.type_of ~refresh:true (Global.env ()) evd f
in
- decompose_prod_n_assum
+ decompose_prod_n_assum evd
(nb_params + nb_args) t,evd
in
let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in
let lemma_type = it_mkProd_or_LetIn eqn type_ctxt in
(* Pp.msgnl (str "lemma type " ++ Printer.pr_lconstr lemma_type ++ fnl () ++ str "f_body " ++ Printer.pr_lconstr f_body); *)
- let f_id = Label.to_id (con_label (fst (destConst f))) in
+ let f_id = Label.to_id (con_label (fst (destConst evd f))) in
let prove_replacement =
tclTHENSEQ
[
@@ -1010,10 +1020,10 @@ let generate_equation_lemma evd fnames f fun_num nb_params nb_args rec_args_num
let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num all_funs g =
let equation_lemma =
try
- let finfos = find_Function_infos (fst (destConst f)) (*FIXME*) in
+ let finfos = find_Function_infos (fst (destConst !evd f)) (*FIXME*) in
mkConst (Option.get finfos.equation_lemma)
with (Not_found | Option.IsNone as e) ->
- let f_id = Label.to_id (con_label (fst (destConst f))) in
+ let f_id = Label.to_id (con_label (fst (destConst !evd f))) in
(*i The next call to mk_equation_id is valid since we will construct the lemma
Ensures by: obvious
i*)
@@ -1022,7 +1032,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
let _ =
match e with
| Option.IsNone ->
- let finfos = find_Function_infos (fst (destConst f)) in
+ let finfos = find_Function_infos (fst (destConst !evd f)) in
update_Function
{finfos with
equation_lemma = Some (match Nametab.locate (qualid_of_ident equation_lemma_id) with
@@ -1038,11 +1048,12 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
(Global.env ()) !evd
(Constrintern.locate_reference (qualid_of_ident equation_lemma_id))
in
+ let res = EConstr.of_constr res in
evd:=evd';
let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd res in
res
in
- let nb_intro_to_do = nb_prod (pf_concl g) in
+ let nb_intro_to_do = nb_prod (project g) (pf_concl g) in
tclTHEN
(tclDO nb_intro_to_do (Proofview.V82.of_tactic intro))
(
@@ -1061,7 +1072,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(* Pp.msgnl (str "princ_type " ++ Printer.pr_lconstr princ_type); *)
(* Pp.msgnl (str "all_funs "); *)
(* Array.iter (fun c -> Pp.msgnl (Printer.pr_lconstr c)) all_funs; *)
- let princ_info = compute_elim_sig princ_type in
+ let princ_info = compute_elim_sig (project g) princ_type in
let fresh_id =
let avoid = ref (pf_ids_of_hyps g) in
(fun na ->
@@ -1090,11 +1101,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
(Global.env ())
(Evd.empty)
- body
+ (EConstr.of_constr body)
| None -> error ( "Cannot define a principle over an axiom ")
in
let fbody = get_body fnames.(fun_num) in
- let f_ctxt,f_body = decompose_lam fbody in
+ let f_ctxt,f_body = decompose_lam (project g) fbody in
let f_ctxt_length = List.length f_ctxt in
let diff_params = princ_info.nparams - f_ctxt_length in
let full_params,princ_params,fbody_with_full_params =
@@ -1129,19 +1140,19 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
princ_params
);
observe (str "fbody_with_full_params := " ++
- pr_lconstr fbody_with_full_params
+ pr_leconstr fbody_with_full_params
);
let all_funs_with_full_params =
Array.map (fun f -> applist(f, List.rev_map var_of_decl full_params)) all_funs
in
let fix_offset = List.length princ_params in
let ptes_to_fix,infos =
- match kind_of_term fbody_with_full_params with
+ match EConstr.kind (project g) fbody_with_full_params with
| Fix((idxs,i),(names,typess,bodies)) ->
let bodies_with_all_params =
Array.map
(fun body ->
- Reductionops.nf_betaiota Evd.empty
+ Reductionops.nf_betaiota (project g)
(applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body,
List.rev_map var_of_decl princ_params))
)
@@ -1150,14 +1161,14 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let info_array =
Array.mapi
(fun i types ->
- let types = prod_applist types (List.rev_map var_of_decl princ_params) in
+ let types = prod_applist (project g) types (List.rev_map var_of_decl princ_params) in
{ idx = idxs.(i) - fix_offset;
name = Nameops.out_name (fresh_id names.(i));
types = types;
offset = fix_offset;
nb_realargs =
List.length
- (fst (decompose_lam bodies.(i))) - fix_offset;
+ (fst (decompose_lam (project g) bodies.(i))) - fix_offset;
body_with_param = bodies_with_all_params.(i);
num_in_block = i
}
@@ -1169,7 +1180,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
(fun i (acc_map,acc_info) decl ->
let pte = RelDecl.get_name decl in
let infos = info_array.(i) in
- let type_args,_ = decompose_prod infos.types in
+ let type_args,_ = decompose_prod (project g) infos.types in
let nargs = List.length type_args in
let f = applist(mkConst fnames.(i), List.rev_map var_of_decl princ_info.params) in
let first_args = Array.init nargs (fun i -> mkRel (nargs -i)) in
@@ -1179,12 +1190,12 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
let body_with_param,num =
let body = get_body fnames.(i) in
let body_with_full_params =
- Reductionops.nf_betaiota Evd.empty (
+ Reductionops.nf_betaiota (project g) (
applist(body,List.rev_map var_of_decl full_params))
in
- match kind_of_term body_with_full_params with
+ match EConstr.kind (project g) body_with_full_params with
| Fix((_,num),(_,_,bs)) ->
- Reductionops.nf_betaiota Evd.empty
+ Reductionops.nf_betaiota (project g)
(
(applist
(substl
@@ -1244,11 +1255,11 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
in
let intros_after_fixes : tactic =
fun gl ->
- let ctxt,pte_app = (decompose_prod_assum (pf_concl gl)) in
- let pte,pte_args = (decompose_app pte_app) in
+ let ctxt,pte_app = (decompose_prod_assum (project gl) (pf_concl gl)) in
+ let pte,pte_args = (decompose_app (project gl) pte_app) in
try
let pte =
- try destVar pte
+ try destVar (project gl) pte
with DestKO -> anomaly (Pp.str "Property is not a variable")
in
let fix_info = Id.Map.find pte ptes_to_fix in
@@ -1267,7 +1278,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
nb_rec_hyps = -100;
rec_hyps = [];
info =
- Reductionops.nf_betaiota Evd.empty
+ Reductionops.nf_betaiota (project g)
(applist(fix_body,List.rev_map mkVar args_id));
eq_hyps = []
}
@@ -1335,7 +1346,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
eq_hyps = []
}
in
- let fname = destConst (fst (decompose_app (List.hd (List.rev pte_args)))) in
+ let fname = destConst (project g) (fst (decompose_app (project g) (List.hd (List.rev pte_args)))) in
tclTHENSEQ
[Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))]);
let do_prove =
@@ -1417,14 +1428,14 @@ let backtrack_eqs_until_hrec hrec eqs : tactic =
let rewrite =
tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) eqs )
in
- let _,hrec_concl = decompose_prod (pf_unsafe_type_of gls (mkVar hrec)) in
- let f_app = Array.last (snd (destApp hrec_concl)) in
- let f = (fst (destApp f_app)) in
+ let _,hrec_concl = decompose_prod (project gls) (pf_unsafe_type_of gls (mkVar hrec)) in
+ let f_app = Array.last (snd (destApp (project gls) hrec_concl)) in
+ let f = (fst (destApp (project gls) f_app)) in
let rec backtrack : tactic =
fun g ->
- let f_app = Array.last (snd (destApp (pf_concl g))) in
- match kind_of_term f_app with
- | App(f',_) when eq_constr f' f -> tclIDTAC g
+ let f_app = Array.last (snd (destApp (project g) (pf_concl g))) in
+ match EConstr.kind (project g) f_app with
+ | App(f',_) when eq_constr (project g) f' f -> tclIDTAC g
| _ -> tclTHEN rewrite backtrack g
in
backtrack gls
@@ -1488,20 +1499,20 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
gls
-let is_valid_hypothesis predicates_name =
+let is_valid_hypothesis sigma predicates_name =
let predicates_name = List.fold_right Id.Set.add predicates_name Id.Set.empty in
let is_pte typ =
- if isApp typ
+ if isApp sigma typ
then
- let pte,_ = destApp typ in
- if isVar pte
- then Id.Set.mem (destVar pte) predicates_name
+ let pte,_ = destApp sigma typ in
+ if isVar sigma pte
+ then Id.Set.mem (destVar sigma pte) predicates_name
else false
else false
in
let rec is_valid_hypothesis typ =
is_pte typ ||
- match kind_of_term typ with
+ match EConstr.kind sigma typ with
| Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ'
| _ -> false
in
@@ -1511,7 +1522,7 @@ let prove_principle_for_gen
(f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes
rec_arg_num rec_arg_type relation gl =
let princ_type = pf_concl gl in
- let princ_info = compute_elim_sig princ_type in
+ let princ_info = compute_elim_sig (project gl) princ_type in
let fresh_id =
let avoid = ref (pf_ids_of_hyps gl) in
fun na ->
@@ -1589,7 +1600,7 @@ let prove_principle_for_gen
let lemma =
match !tcc_lemma_ref with
| None -> error "No tcc proof !!"
- | Some lemma -> lemma
+ | Some lemma -> EConstr.of_constr lemma
in
(* let rec list_diff del_list check_list = *)
(* match del_list with *)
@@ -1649,7 +1660,7 @@ let prove_principle_for_gen
Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref));
(* observe_tac "finish" *) (fun gl' ->
let body =
- let _,args = destApp (pf_concl gl') in
+ let _,args = destApp (project gl') (pf_concl gl') in
Array.last args
in
let body_info rec_hyps =
@@ -1692,7 +1703,7 @@ let prove_principle_for_gen
)
);
- is_valid = is_valid_hypothesis predicates_names
+ is_valid = is_valid_hypothesis (project gl') predicates_names
}
in
let ptes_info : pte_info Id.Map.t =
diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli
index 34ce66967..769d726d7 100644
--- a/plugins/funind/functional_principles_proofs.mli
+++ b/plugins/funind/functional_principles_proofs.mli
@@ -4,7 +4,7 @@ open Term
val prove_princ_for_struct :
Evd.evar_map ref ->
bool ->
- int -> constant array -> constr array -> int -> Tacmach.tactic
+ int -> constant array -> EConstr.constr array -> int -> Tacmach.tactic
val prove_principle_for_gen :
@@ -12,8 +12,8 @@ val prove_principle_for_gen :
constr option ref -> (* a pointer to the obligation proofs lemma *)
bool -> (* is that function uses measure *)
int -> (* the number of recursive argument *)
- types -> (* the type of the recursive argument *)
- constr -> (* the wf relation used to prove the function *)
+ EConstr.types -> (* the type of the recursive argument *)
+ EConstr.constr -> (* the wf relation used to prove the function *)
Tacmach.tactic
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index cc699e5d3..529b91c4c 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -23,16 +23,19 @@ let observe s =
if do_observe ()
then Feedback.msg_debug s
+let pop t = Vars.lift (-1) t
+
(*
Transform an inductive induction principle into
a functional one
*)
let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
- let princ_type_info = compute_elim_sig princ_type in
+ let princ_type = EConstr.of_constr princ_type in
+ let princ_type_info = compute_elim_sig Evd.empty princ_type (** FIXME *) in
let env = Global.env () in
- let env_with_params = Environ.push_rel_context princ_type_info.params env in
+ let env_with_params = EConstr.push_rel_context princ_type_info.params env in
let tbl = Hashtbl.create 792 in
- let rec change_predicates_names (avoid:Id.t list) (predicates:Context.Rel.t) : Context.Rel.t =
+ let rec change_predicates_names (avoid:Id.t list) (predicates:EConstr.rel_context) : EConstr.rel_context =
match predicates with
| [] -> []
| decl :: predicates ->
@@ -53,14 +56,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
(* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *)
let change_predicate_sort i decl =
let new_sort = sorts.(i) in
- let args,_ = decompose_prod (RelDecl.get_type decl) in
+ let args,_ = decompose_prod (EConstr.Unsafe.to_constr (RelDecl.get_type decl)) in
let real_args =
if princ_type_info.indarg_in_concl
then List.tl args
else args
in
Context.Named.Declaration.LocalAssum (Nameops.out_name (Context.Rel.Declaration.get_name decl),
- compose_prod real_args (mkSort new_sort))
+ Term.compose_prod real_args (mkSort new_sort))
in
let new_predicates =
List.map_i
@@ -84,6 +87,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| _ -> false
in
let pre_princ =
+ let open EConstr in
it_mkProd_or_LetIn
(it_mkProd_or_LetIn
(Option.fold_right
@@ -95,6 +99,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
)
princ_type_info.branches
in
+ 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
@@ -110,7 +115,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
let dummy_var = mkVar (Id.of_string "________") in
let mk_replacement c i args =
- let res = mkApp(rel_to_fun.(i), Array.map Termops.pop (array_get_start args)) in
+ let res = mkApp(rel_to_fun.(i), Array.map pop (array_get_start args)) in
observe (str "replacing " ++ pr_lconstr c ++ str " by " ++ pr_lconstr res);
res
in
@@ -168,25 +173,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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 (Termops.pop new_b), filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
+ then (pop new_b), filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b
else
(
bind_fun(new_x,new_t,new_b),
list_union_eq
eq_constr
binders_to_remove_from_t
- (List.map Termops.pop binders_to_remove_from_b)
+ (List.map pop binders_to_remove_from_b)
)
with
| Toberemoved ->
(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *)
let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in
- new_b, List.map Termops.pop binders_to_remove_from_b
+ new_b, List.map pop binders_to_remove_from_b
| Toberemoved_with_rel (n,c) ->
(* observe (str "Decl of "++Ppconstr.pr_name 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 Termops.pop binders_to_remove_from_b)
+ new_b, list_add_set_eq eq_constr (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
@@ -197,25 +202,25 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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 (Termops.pop new_b),filter_map (eq_constr (mkRel 1)) Termops.pop binders_to_remove_from_b
+ then (pop new_b),filter_map (eq_constr (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)
- (List.map Termops.pop binders_to_remove_from_b)
+ (List.map pop binders_to_remove_from_b)
)
with
| Toberemoved ->
(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *)
let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in
- new_b, List.map Termops.pop binders_to_remove_from_b
+ new_b, List.map pop binders_to_remove_from_b
| Toberemoved_with_rel (n,c) ->
(* observe (str "Decl of "++Ppconstr.pr_name 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 Termops.pop binders_to_remove_from_b)
+ new_b, list_add_set_eq eq_constr (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
@@ -237,20 +242,21 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
| Context.Named.Declaration.LocalDef (id,t,b) -> LocalDef (Name (Hashtbl.find tbl id), t, b))
new_predicates)
)
- princ_type_info.params
+ (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_type_info.params)
let change_property_sort evd toSort princ princName =
let open Context.Rel.Declaration in
- let princ_info = compute_elim_sig princ in
+ let princ = EConstr.of_constr princ in
+ let princ_info = compute_elim_sig evd princ in
let change_sort_in_predicate decl =
LocalAssum
(get_name decl,
- let args,ty = decompose_prod (get_type decl) in
+ let args,ty = decompose_prod (EConstr.Unsafe.to_constr (get_type decl)) in
let s = destSort ty in
Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty);
- compose_prod args (mkSort toSort)
+ Term.compose_prod args (mkSort toSort)
)
in
let evd,princName_as_constr =
@@ -266,11 +272,11 @@ let change_property_sort evd toSort princ princName =
(it_mkLambda_or_LetIn init
(List.map change_sort_in_predicate princ_info.predicates)
)
- princ_info.params
+ (List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) princ_info.params)
let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_princ_type sorts funs i proof_tac hook =
(* First we get the type of the old graph principle *)
- let mutr_nparams = (compute_elim_sig old_princ_type).nparams in
+ let mutr_nparams = (compute_elim_sig !evd (EConstr.of_constr old_princ_type)).nparams in
(* let time1 = System.get_time () in *)
let new_principle_type =
compute_new_princ_type_from_rel
@@ -283,18 +289,19 @@ let build_functional_principle (evd:Evd.evar_map ref) interactive_proof old_prin
let new_princ_name =
next_ident_away_in_goal (Id.of_string "___________princ_________") []
in
- let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd new_principle_type in
+ let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr new_principle_type) in
let hook = Lemmas.mk_hook (hook new_principle_type) in
begin
Lemmas.start_proof
new_princ_name
(Decl_kinds.Global,Flags.is_universe_polymorphism (),(Decl_kinds.Proof Decl_kinds.Theorem))
!evd
- new_principle_type
+ (EConstr.of_constr new_principle_type)
hook
;
(* let _tim1 = System.get_time () in *)
- ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map mkConstU funs) mutr_nparams)));
+ let map (c, u) = EConstr.mkConstU (c, EConstr.EInstance.make u) in
+ ignore (Pfedit.by (Proofview.V82.tactic (proof_tac (Array.map map funs) mutr_nparams)));
(* let _tim2 = System.get_time () in *)
(* begin *)
(* let dur1 = System.time_difference tim1 tim2 in *)
@@ -337,7 +344,7 @@ let generate_functional_principle (evd: Evd.evar_map ref)
let evd',s = Evd.fresh_sort_in_family env evd' fam_sort in
let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in
let evd',value = change_property_sort evd' s new_principle_type new_princ_name in
- let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' value) in
+ let evd' = fst (Typing.type_of ~refresh:true (Global.env ()) evd' (EConstr.of_constr value)) in
(* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *)
let ce = Declare.definition_entry ~poly:(Flags.is_universe_polymorphism ()) ~univs:(snd (Evd.universe_context evd')) value in
ignore(
@@ -405,8 +412,9 @@ let get_funs_constant mp dp =
(CClosure.RedFlags.mkflags [CClosure.RedFlags.fZETA])
(Global.env ())
(Evd.from_env (Global.env ()))
- body
+ (EConstr.of_constr body)
in
+ let body = EConstr.Unsafe.to_constr body in
body
| None -> error ( "Cannot define a principle over an axiom ")
in
@@ -488,7 +496,7 @@ let make_scheme evd (fas : (pconstant*glob_sort) list) : Safe_typing.private_con
in
let _ = evd := sigma in
let l_schemes =
- List.map (Typing.unsafe_type_of env sigma) schemes
+ List.map (EConstr.of_constr %> Typing.unsafe_type_of env sigma %> EConstr.Unsafe.to_constr) schemes
in
let i = ref (-1) in
let sorts =
@@ -616,7 +624,7 @@ let build_scheme fas =
in
let evd',f = Evd.fresh_global (Global.env ()) !evd f_as_constant in
let _ = evd := evd' in
- let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd f in
+ let _ = Typing.e_type_of ~refresh:true (Global.env ()) evd (EConstr.of_constr f) in
(destConst f,sort)
)
fas
@@ -666,7 +674,7 @@ let build_case_scheme fa =
Indrec.build_case_analysis_scheme_default env sigma ind sf
in
let sigma = Sigma.to_evar_map sigma in
- let scheme_type = (Typing.unsafe_type_of env sigma ) scheme in
+ let scheme_type = EConstr.Unsafe.to_constr ((Typing.unsafe_type_of env sigma) (EConstr.of_constr scheme)) in
let sorts =
(fun (_,_,x) ->
Universes.new_sort_in_family (Pretyping.interp_elimination_sort x)
diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli
index 3fa2644ca..45ad332fc 100644
--- a/plugins/funind/functional_principles_types.mli
+++ b/plugins/funind/functional_principles_types.mli
@@ -27,7 +27,7 @@ val generate_functional_principle :
(* The tactic to use to make the proof w.r
the number of params
*)
- (constr array -> int -> Tacmach.tactic) ->
+ (EConstr.constr array -> int -> Tacmach.tactic) ->
unit
val compute_new_princ_type_from_rel : constr array -> sorts array ->
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index cf2e42d2c..0dccd25d7 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -98,7 +98,8 @@ ARGUMENT EXTEND with_names TYPED AS intropattern_opt PRINTED BY pr_intro_as_pat
| [] ->[ None ]
END
-
+let functional_induction b c x pat =
+ Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat))
TACTIC EXTEND newfunind
@@ -107,9 +108,9 @@ TACTIC EXTEND newfunind
let c = match cl with
| [] -> assert false
| [c] -> c
- | c::cl -> applist(c,cl)
+ | c::cl -> EConstr.applist(c,cl)
in
- Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat))) princl ]
+ Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ]
END
(***** debug only ***)
TACTIC EXTEND snewfunind
@@ -118,9 +119,9 @@ TACTIC EXTEND snewfunind
let c = match cl with
| [] -> assert false
| [c] -> c
- | c::cl -> applist(c,cl)
+ | c::cl -> EConstr.applist(c,cl)
in
- Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction false c x (Option.map out_disjunctive pat))) princl ]
+ Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ]
END
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 084de31c0..7dc869131 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -352,7 +352,7 @@ let add_pat_variables pat typ env : Environ.env =
| PatVar(_,na) -> Environ.push_rel (RelDecl.LocalAssum (na,typ)) env
| PatCstr(_,c,patl,na) ->
let Inductiveops.IndType(indf,indargs) =
- try Inductiveops.find_rectype env (Evd.from_env env) typ
+ try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ)
with Not_found -> assert false
in
let constructors = Inductiveops.get_constructors env indf in
@@ -409,7 +409,7 @@ let rec pattern_to_term_and_type env typ = function
constr
in
let Inductiveops.IndType(indf,indargs) =
- try Inductiveops.find_rectype env (Evd.from_env env) typ
+ try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ)
with Not_found -> assert false
in
let constructors = Inductiveops.get_constructors env indf in
@@ -421,7 +421,7 @@ let rec pattern_to_term_and_type env typ = function
Array.to_list
(Array.init
(cst_narg - List.length patternl)
- (fun i -> Detyping.detype false [] env (Evd.from_env env) csta.(i))
+ (fun i -> Detyping.detype false [] env (Evd.from_env env) (EConstr.of_constr csta.(i)))
)
in
let patl_as_term =
@@ -503,7 +503,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
The "value" of this branch is then simply [res]
*)
let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in
- let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in
+ let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr rt_as_constr) in
let res_raw_type = Detyping.detype false [] env (Evd.from_env env) rt_typ in
let res = fresh_id args_res.to_avoid "_res" in
let new_avoid = res::args_res.to_avoid in
@@ -612,7 +612,8 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let v = match typ with None -> v | Some t -> GCast (loc,v,CastConv t) in
let v_res = build_entry_lc env funnames avoid v in
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
- let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in
+ let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in
+ let v_type = EConstr.Unsafe.to_constr v_type in
let new_env =
match n with
Anonymous -> env
@@ -628,7 +629,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
build_entry_lc_from_case env funnames make_discr el brl avoid
| GIf(_,b,(na,e_option),lhs,rhs) ->
let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
- let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in
let (ind,_) =
try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
@@ -660,7 +661,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
nal
in
let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
- let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in
+ let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in
let (ind,_) =
try Inductiveops.find_inductive env (Evd.from_env env) b_typ
with Not_found ->
@@ -707,7 +708,7 @@ and build_entry_lc_from_case env funname make_discr
let types =
List.map (fun (case_arg,_) ->
let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in
- Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr
+ EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr case_arg_as_constr))
) el
in
(****** The next works only if the match is not dependent ****)
@@ -754,7 +755,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
List.fold_right
(fun id acc ->
let typ_of_id =
- Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (mkVar id)
+ Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id)
in
let raw_typ_of_id =
Detyping.detype false []
@@ -802,13 +803,14 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
List.map3
(fun pat e typ_as_constr ->
let this_pat_ids = ids_of_pat pat in
+ let typ_as_constr = EConstr.of_constr typ_as_constr in
let typ = Detyping.detype false [] new_env (Evd.from_env env) typ_as_constr in
let pat_as_term = pattern_to_term pat in
List.fold_right
(fun id acc ->
if Id.Set.mem id this_pat_ids
then (Prod (Name id),
- let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (mkVar id) in
+ let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in
let raw_typ_of_id =
Detyping.detype false [] new_env (Evd.from_env env) typ_of_id
in
@@ -953,7 +955,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
mkGProd(n,t,new_b),id_to_exclude
with Continue ->
- let jmeq = Globnames.IndRef (fst (destInd (jmeq ()))) in
+ let jmeq = Globnames.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in
let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in
let ind,args' = Inductive.find_inductive env ty' in
let mib,_ = Global.lookup_inductive (fst ind) in
@@ -967,7 +969,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(List.map
(fun p -> Detyping.detype false []
env (Evd.from_env env)
- p) params)@(Array.to_list
+ (EConstr.of_constr p)) params)@(Array.to_list
(Array.make
(List.length args' - nparam)
(mkGHole ()))))
@@ -985,6 +987,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let ty' = snd (Util.List.chop nparam ty) in
List.fold_left2
(fun acc var_as_constr arg ->
+ let arg = EConstr.of_constr arg in
if isRel var_as_constr
then
let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in
@@ -1123,7 +1126,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let evd = (Evd.from_env env) in
let t',ctx = Pretyping.understand env evd t in
let evd = Evd.from_ctx ctx in
- let type_t' = Typing.unsafe_type_of env evd t' in
+ let type_t' = Typing.unsafe_type_of env evd (EConstr.of_constr t') in
+ let type_t' = EConstr.Unsafe.to_constr type_t' in
let new_env = Environ.push_rel (LocalDef (n,t',type_t')) env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1277,8 +1281,10 @@ let do_build_inductive
let open Context.Named.Declaration in
let evd,env =
Array.fold_right2
- (fun id c (evd,env) ->
- let evd,t = Typing.type_of env evd (mkConstU c) in
+ (fun id (c, u) (evd,env) ->
+ let u = EConstr.EInstance.make u in
+ let evd,t = Typing.type_of env evd (EConstr.mkConstU (c, u)) in
+ let t = EConstr.Unsafe.to_constr t in
evd,
Environ.push_named (LocalAssum (id,t))
(* try *)
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index d394fe313..ebeddf5f6 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -2,6 +2,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Pp
open Indfun_common
open Libnames
@@ -14,37 +15,39 @@ open Sigma.Notations
module RelDecl = Context.Rel.Declaration
-let is_rec_info scheme_info =
+let is_rec_info sigma scheme_info =
let test_branche min acc decl =
acc || (
let new_branche =
- it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum (RelDecl.get_type decl))) in
- let free_rels_in_br = Termops.free_rels new_branche in
+ it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum sigma (RelDecl.get_type decl))) in
+ let free_rels_in_br = Termops.free_rels sigma new_branche in
let max = min + scheme_info.Tactics.npredicates in
Int.Set.exists (fun i -> i >= min && i< max) free_rels_in_br
)
in
List.fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches)
-let choose_dest_or_ind scheme_info =
- Tactics.induction_destruct (is_rec_info scheme_info) false
+let choose_dest_or_ind scheme_info args =
+ Proofview.tclBIND Proofview.tclEVARMAP (fun sigma ->
+ Tactics.induction_destruct (is_rec_info sigma scheme_info) false args)
let functional_induction with_clean c princl pat =
let res =
- let f,args = decompose_app c in
fun g ->
+ let sigma = Tacmach.project g in
+ let f,args = decompose_app sigma c in
let princ,bindings, princ_type,g' =
match princl with
| None -> (* No principle is given let's find the good one *)
begin
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Const (c',u) ->
let princ_option =
let finfo = (* we first try to find out a graph on f *)
try find_Function_infos c'
with Not_found ->
user_err (str "Cannot find induction information on "++
- Printer.pr_lconstr (mkConst c') )
+ Printer.pr_leconstr (mkConst c') )
in
match Tacticals.elimination_sort_of_goal g with
| InProp -> finfo.prop_lemma
@@ -72,15 +75,17 @@ let functional_induction with_clean c princl pat =
(* mkConst(const_of_id princ_name ),g (\* FIXME *\) *)
with Not_found -> (* This one is neither defined ! *)
user_err (str "Cannot find induction principle for "
- ++Printer.pr_lconstr (mkConst c') )
+ ++Printer.pr_leconstr (mkConst c') )
in
- (princ,NoBindings, Tacmach.pf_unsafe_type_of g' princ,g')
+ let princ = EConstr.of_constr princ in
+ (princ,NoBindings,Tacmach.pf_unsafe_type_of g' princ,g')
| _ -> raise (UserError(None,str "functional induction must be used with a function" ))
end
| Some ((princ,binding)) ->
princ,binding,Tacmach.pf_unsafe_type_of g princ,g
in
- let princ_infos = Tactics.compute_elim_sig princ_type in
+ let sigma = Tacmach.project g' in
+ let princ_infos = Tactics.compute_elim_sig (Tacmach.project g') princ_type in
let args_as_induction_constr =
let c_list =
if princ_infos.Tactics.farg_in_concl
@@ -94,7 +99,7 @@ let functional_induction with_clean c princl pat =
let princ' = Some (princ,bindings) in
let princ_vars =
List.fold_right
- (fun a acc -> try Id.Set.add (destVar a) acc with DestKO -> acc)
+ (fun a acc -> try Id.Set.add (destVar sigma a) acc with DestKO -> acc)
args
Id.Set.empty
in
@@ -245,7 +250,9 @@ let derive_inversion fix_names =
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident id)) in
- evd, destConst c::l
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ evd, (cst, EInstance.kind evd u) :: l
)
fix_names
(evd',[])
@@ -265,7 +272,8 @@ let derive_inversion fix_names =
(Global.env ()) evd
(Constrintern.locate_reference (Libnames.qualid_of_ident (mk_rel_id id)))
in
- evd,(fst (destInd id))::l
+ let id = EConstr.of_constr id in
+ evd,(fst (destInd evd id))::l
)
fix_names
(evd',[])
@@ -332,7 +340,7 @@ let error_error names e =
let generate_principle (evd:Evd.evar_map ref) pconstants on_error
is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof
- (continue_proof : int -> Names.constant array -> Term.constr array -> int ->
+ (continue_proof : int -> Names.constant array -> EConstr.constr array -> int ->
Tacmach.tactic) : unit =
let names = List.map (function (((_, name),_),_,_,_,_),_ -> name) fix_rec_l in
let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in
@@ -370,7 +378,8 @@ let generate_principle (evd:Evd.evar_map ref) pconstants on_error
let evd = ref (Evd.from_env env) in
let evd',uprinc = Evd.fresh_global env !evd princ in
let _ = evd := evd' in
- let princ_type = Typing.e_type_of ~refresh:true env evd uprinc in
+ let princ_type = Typing.e_type_of ~refresh:true env evd (EConstr.of_constr uprinc) in
+ let princ_type = EConstr.Unsafe.to_constr princ_type in
Functional_principles_types.generate_functional_principle
evd
interactive_proof
@@ -405,7 +414,10 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- evd,((destConst c)::l)
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ let u = EInstance.kind evd u in
+ evd,((cst, u) :: l)
)
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
@@ -419,7 +431,10 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let evd,c =
Evd.fresh_global
(Global.env ()) evd (Constrintern.locate_reference (Libnames.qualid_of_ident fname)) in
- evd,((destConst c)::l)
+ let c = EConstr.of_constr c in
+ let (cst, u) = destConst evd c in
+ let u = EInstance.kind evd u in
+ evd,((cst, u) :: l)
)
(Evd.from_env (Global.env ()),[])
fixpoint_exprl
@@ -429,7 +444,7 @@ let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexp
let generate_correction_proof_wf f_ref tcc_lemma_ref
is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation
- (_: int) (_:Names.constant array) (_:Term.constr array) (_:int) : Tacmach.tactic =
+ (_: int) (_:Names.constant array) (_:EConstr.constr array) (_:int) : Tacmach.tactic =
Functional_principles_proofs.prove_principle_for_gen
(f_ref,functional_ref,eq_ref)
tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation
@@ -837,7 +852,7 @@ let make_graph (f_ref:global_reference) =
| ConstRef c ->
begin try c,Global.lookup_constant c
with Not_found ->
- raise (UserError (None,str "Cannot find " ++ Printer.pr_lconstr (mkConst c)) )
+ raise (UserError (None,str "Cannot find " ++ Printer.pr_leconstr (mkConst c)) )
end
| _ -> raise (UserError (None, str "Not a function reference") )
in
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index 1c27bdfac..ba89fe4a7 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -12,8 +12,8 @@ val do_generate_principle :
val functional_induction :
bool ->
- Term.constr ->
- (Term.constr * Term.constr bindings) option ->
+ EConstr.constr ->
+ (EConstr.constr * EConstr.constr bindings) option ->
Tacexpr.or_and_intro_pattern option ->
Proof_type.goal Tacmach.sigma -> Proof_type.goal list Evd.sigma
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index aed0fa331..20da12f39 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -130,8 +130,8 @@ let find_reference sl s =
let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in
Nametab.locate (make_qualid dp (Id.of_string s))
-let eq = lazy(coq_constant "eq")
-let refl_equal = lazy(coq_constant "eq_refl")
+let eq = lazy(EConstr.of_constr (coq_constant "eq"))
+let refl_equal = lazy(EConstr.of_constr (coq_constant "eq_refl"))
(*****************************************************************)
(* Copy of the standart save mechanism but without the much too *)
@@ -475,13 +475,13 @@ exception ToShow of exn
let jmeq () =
try
Coqlib.check_required_library Coqlib.jmeq_module_name;
- Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq"
+ EConstr.of_constr (Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq")
with e when CErrors.noncritical e -> raise (ToShow e)
let jmeq_refl () =
try
Coqlib.check_required_library Coqlib.jmeq_module_name;
- Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl"
+ EConstr.of_constr (Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl")
with e when CErrors.noncritical e -> raise (ToShow e)
let h_intros l =
@@ -489,10 +489,10 @@ let h_intros l =
let h_id = Id.of_string "h"
let hrec_id = Id.of_string "hrec"
-let well_founded = function () -> (coq_constant "well_founded")
-let acc_rel = function () -> (coq_constant "Acc")
-let acc_inv_id = function () -> (coq_constant "Acc_inv")
-let well_founded_ltof = function () -> (Coqlib.coq_constant "" ["Arith";"Wf_nat"] "well_founded_ltof")
+let well_founded = function () -> EConstr.of_constr (coq_constant "well_founded")
+let acc_rel = function () -> EConstr.of_constr (coq_constant "Acc")
+let acc_inv_id = function () -> EConstr.of_constr (coq_constant "Acc_inv")
+let well_founded_ltof = function () -> EConstr.of_constr (Coqlib.coq_constant "" ["Arith";"Wf_nat"] "well_founded_ltof")
let ltof_ref = function () -> (find_reference ["Coq";"Arith";"Wf_nat"] "ltof")
let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (Global.env ()) *)
@@ -501,8 +501,45 @@ let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (G
| VarRef id -> EvalVarRef id
| _ -> assert false;;
-let list_rewrite (rev:bool) (eqs: (constr*bool) list) =
+let list_rewrite (rev:bool) (eqs: (EConstr.constr*bool) list) =
tclREPEAT
(List.fold_right
(fun (eq,b) i -> tclORELSE (Proofview.V82.of_tactic ((if b then Equality.rewriteLR else Equality.rewriteRL) eq)) i)
(if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));;
+
+let decompose_lam_n sigma n =
+ let open EConstr in
+ if n < 0 then CErrors.error "decompose_lam_n: integer parameter must be positive";
+ let rec lamdec_rec l n c =
+ if Int.equal n 0 then l,c
+ else match EConstr.kind sigma c with
+ | Lambda (x,t,c) -> lamdec_rec ((x,t)::l) (n-1) c
+ | Cast (c,_,_) -> lamdec_rec l n c
+ | _ -> CErrors.error "decompose_lam_n: not enough abstractions"
+ in
+ lamdec_rec [] n
+
+let lamn n env b =
+ let open EConstr in
+ let rec lamrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> lamrec (n-1, l, mkLambda (v,t,b))
+ | _ -> assert false
+ in
+ lamrec (n,env,b)
+
+(* compose_lam [xn:Tn;..;x1:T1] b = [x1:T1]..[xn:Tn]b *)
+let compose_lam l b = lamn (List.length l) l b
+
+(* prodn n [xn:Tn;..;x1:T1;Gamma] b = (x1:T1)..(xn:Tn)b *)
+let prodn n env b =
+ let open EConstr in
+ let rec prodrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b))
+ | _ -> assert false
+ in
+ prodrec (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
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index 2aabfa003..5c3e73e9d 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -40,11 +40,11 @@ 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 eq : Term.constr Lazy.t
-val refl_equal : Term.constr Lazy.t
+val eq : EConstr.constr Lazy.t
+val refl_equal : EConstr.constr Lazy.t
val const_of_id: Id.t -> Globnames.global_reference(* constantyes *)
-val jmeq : unit -> Term.constr
-val jmeq_refl : unit -> Term.constr
+val jmeq : unit -> EConstr.constr
+val jmeq_refl : unit -> EConstr.constr
val save : bool -> Id.t -> Safe_typing.private_constants Entries.definition_entry -> Decl_kinds.goal_kind ->
unit Lemmas.declaration_hook CEphemeron.key -> unit
@@ -107,10 +107,15 @@ val is_strict_tcc : unit -> bool
val h_intros: Names.Id.t list -> Proof_type.tactic
val h_id : Names.Id.t
val hrec_id : Names.Id.t
-val acc_inv_id : Term.constr Util.delayed
+val acc_inv_id : EConstr.constr Util.delayed
val ltof_ref : Globnames.global_reference Util.delayed
-val well_founded_ltof : Term.constr Util.delayed
-val acc_rel : Term.constr Util.delayed
-val well_founded : Term.constr Util.delayed
+val well_founded_ltof : EConstr.constr Util.delayed
+val acc_rel : EConstr.constr Util.delayed
+val well_founded : EConstr.constr Util.delayed
val evaluable_of_global_reference : Globnames.global_reference -> Names.evaluable_global_reference
-val list_rewrite : bool -> (Term.constr*bool) list -> Proof_type.tactic
+val list_rewrite : bool -> (EConstr.constr*bool) list -> Proof_type.tactic
+
+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
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 70333b063..94ec0a898 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -13,6 +13,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Vars
open Pp
open Globnames
@@ -109,11 +110,11 @@ let thin ids gl = Proofview.V82.of_tactic (Tactics.clear ids) gl
let make_eq () =
try
- Universes.constr_of_global (Coqlib.build_coq_eq ())
+ EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ()))
with _ -> assert false
let make_eq_refl () =
try
- Universes.constr_of_global (Coqlib.build_coq_eq_refl ())
+ EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq_refl ()))
with _ -> assert false
@@ -132,11 +133,12 @@ let make_eq_refl () =
let generate_type evd g_to_f f graph i =
(*i we deduce the number of arguments of the function and its returned type from the graph i*)
let evd',graph =
- Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd graph)))
+ Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd !evd graph)))
in
+ let graph = EConstr.of_constr graph in
evd:=evd';
let graph_arity = Typing.e_type_of (Global.env ()) evd graph in
- let ctxt,_ = decompose_prod_assum graph_arity in
+ let ctxt,_ = decompose_prod_assum !evd graph_arity in
let fun_ctxt,res_type =
match ctxt with
| [] | [_] -> anomaly (Pp.str "Not a valid context")
@@ -194,7 +196,7 @@ let generate_type evd g_to_f f graph i =
WARNING: while convertible, [type_of body] and [type] can be non equal
*)
let find_induction_principle evd f =
- let f_as_constant,u = match kind_of_term f with
+ let f_as_constant,u = match EConstr.kind !evd f with
| Const c' -> c'
| _ -> error "Must be used with a function"
in
@@ -203,6 +205,7 @@ let find_induction_principle evd f =
| None -> raise Not_found
| Some rect_lemma ->
let evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in
+ let rect_lemma = EConstr.of_constr rect_lemma in
let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in
evd:=evd';
rect_lemma,typ
@@ -247,15 +250,15 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
\[fun (x_1:t_1)\ldots(x_n:t_n)=> fun fv => fun res => res = fv \rightarrow graph\ x_1\ldots x_n\ res\]
*)
(* we the get the definition of the graphs block *)
- let graph_ind,u = destInd graphs_constr.(i) in
+ let graph_ind,u = destInd evd graphs_constr.(i) in
let kn = fst graph_ind in
let mib,_ = Global.lookup_inductive graph_ind in
(* and the principle to use in this lemma in $\zeta$ normal form *)
let f_principle,princ_type = schemes.(i) in
let princ_type = nf_zeta princ_type in
- let princ_infos = Tactics.compute_elim_sig princ_type in
+ let princ_infos = Tactics.compute_elim_sig evd princ_type in
(* The number of args of the function is then easily computable *)
- let nb_fun_args = nb_prod (pf_concl g) - 2 in
+ let nb_fun_args = nb_prod (project g) (pf_concl g) - 2 in
let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in
let ids = args_names@(pf_ids_of_hyps g) in
(* Since we cannot ensure that the functional principle is defined in the
@@ -272,13 +275,13 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(fun decl ->
List.map
(fun id -> Loc.ghost, IntroNaming (IntroIdentifier id))
- (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum (RelDecl.get_type decl)))))
+ (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl)))))
)
branches
in
(* before building the full intro pattern for the principle *)
let eq_ind = make_eq () in
- let eq_construct = mkConstructUi (destInd eq_ind, 1) in
+ let eq_construct = mkConstructUi (destInd evd eq_ind, 1) in
(* The next to referencies will be used to find out which constructor to apply in each branch *)
let ind_number = ref 0
and min_constr_number = ref 0 in
@@ -307,17 +310,18 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
List.fold_right
(fun hid acc ->
let type_of_hid = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term type_of_hid with
+ let sigma = project g in
+ match EConstr.kind sigma type_of_hid with
| Prod(_,_,t') ->
begin
- match kind_of_term t' with
+ match EConstr.kind sigma t' with
| Prod(_,t'',t''') ->
begin
- match kind_of_term t'',kind_of_term t''' with
+ match EConstr.kind sigma t'',EConstr.kind sigma t''' with
| App(eq,args), App(graph',_)
when
- (eq_constr eq eq_ind) &&
- Array.exists (Constr.eq_constr_nounivs graph') graphs_constr ->
+ (EConstr.eq_constr sigma eq eq_ind) &&
+ Array.exists (EConstr.eq_constr_nounivs sigma graph') graphs_constr ->
(args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|]))
::acc)
| _ -> mkVar hid :: acc
@@ -400,8 +404,8 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
match ctxt with
| [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
| hres::res::decl::ctxt ->
- let res = Termops.it_mkLambda_or_LetIn
- (Termops.it_mkProd_or_LetIn concl [hres;res])
+ let res = EConstr.it_mkLambda_or_LetIn
+ (EConstr.it_mkProd_or_LetIn concl [hres;res])
(LocalAssum (RelDecl.get_name decl, RelDecl.get_type decl) :: ctxt)
in
res
@@ -468,7 +472,7 @@ let generalize_dependent_of x hyp g =
tclMAP
(function
| LocalAssum (id,t) when not (Id.equal id hyp) &&
- (Termops.occur_var (pf_env g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
+ (Termops.occur_var (pf_env g) (project g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
| _ -> tclIDTAC
)
(pf_hyps g)
@@ -492,43 +496,44 @@ let rec intros_with_rewrite g =
and intros_with_rewrite_aux : tactic =
fun g ->
let eq_ind = make_eq () in
- match kind_of_term (pf_concl g) with
+ let sigma = project g in
+ match EConstr.kind sigma (pf_concl g) with
| Prod(_,t,t') ->
begin
- match kind_of_term t with
- | App(eq,args) when (eq_constr eq eq_ind) ->
+ match EConstr.kind sigma t with
+ | App(eq,args) when (EConstr.eq_constr sigma eq eq_ind) ->
if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2)
then
let id = pf_get_new_id (Id.of_string "y") g in
tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g
- else if isVar args.(1) && (Environ.evaluable_named (destVar args.(1)) (pf_env g))
+ else if isVar sigma args.(1) && (Environ.evaluable_named (destVar sigma args.(1)) (pf_env g))
then tclTHENSEQ[
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))]);
- tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar args.(1)),Locus.InHyp) )))
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))] ((destVar sigma args.(1)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
- else if isVar args.(2) && (Environ.evaluable_named (destVar args.(2)) (pf_env g))
+ else if isVar sigma args.(2) && (Environ.evaluable_named (destVar sigma args.(2)) (pf_env g))
then tclTHENSEQ[
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))]);
- tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) )))
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))] ((destVar sigma args.(2)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
- else if isVar args.(1)
+ else if isVar sigma args.(1)
then
let id = pf_get_new_id (Id.of_string "y") g in
tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);
- generalize_dependent_of (destVar args.(1)) id;
+ generalize_dependent_of (destVar sigma args.(1)) id;
tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id)));
intros_with_rewrite
]
g
- else if isVar args.(2)
+ else if isVar sigma args.(2)
then
let id = pf_get_new_id (Id.of_string "y") g in
tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);
- generalize_dependent_of (destVar args.(2)) id;
+ generalize_dependent_of (destVar sigma args.(2)) id;
tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id)));
intros_with_rewrite
]
@@ -542,7 +547,7 @@ and intros_with_rewrite_aux : tactic =
intros_with_rewrite
] g
end
- | Ind _ when eq_constr t (Coqlib.build_coq_False ()) ->
+ | Ind _ when EConstr.eq_constr sigma t (EConstr.of_constr (Coqlib.build_coq_False ())) ->
Proofview.V82.of_tactic tauto g
| Case(_,_,v,_) ->
tclTHENSEQ[
@@ -580,7 +585,7 @@ and intros_with_rewrite_aux : tactic =
let rec reflexivity_with_destruct_cases g =
let destruct_case () =
try
- match kind_of_term (snd (destApp (pf_concl g))).(2) with
+ match EConstr.kind (project g) (snd (destApp (project g) (pf_concl g))).(2) with
| Case(_,_,v,_) ->
tclTHENSEQ[
Proofview.V82.of_tactic (simplest_case v);
@@ -597,8 +602,8 @@ let rec reflexivity_with_destruct_cases g =
match sc with
None -> tclIDTAC g
| Some id ->
- match kind_of_term (pf_unsafe_type_of g (mkVar id)) with
- | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind ->
+ match EConstr.kind (project g) (pf_unsafe_type_of g (mkVar id)) with
+ | App(eq,[|_;t1;t2|]) when EConstr.eq_constr (project g) eq eq_ind ->
if Equality.discriminable (pf_env g) (project g) t1 t2
then Proofview.V82.of_tactic (Equality.discrHyp id) g
else if Equality.injectable (pf_env g) (project g) t1 t2
@@ -656,18 +661,18 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
*)
let lemmas =
Array.map
- (fun (_,(ctxt,concl)) -> nf_zeta (Termops.it_mkLambda_or_LetIn concl ctxt))
+ (fun (_,(ctxt,concl)) -> nf_zeta (EConstr.it_mkLambda_or_LetIn concl ctxt))
lemmas_types_infos
in
(* We get the constant and the principle corresponding to this lemma *)
let f = funcs.(i) in
- let graph_principle = nf_zeta schemes.(i) in
+ let graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in
let princ_type = pf_unsafe_type_of g graph_principle in
- let princ_infos = Tactics.compute_elim_sig princ_type in
+ let princ_infos = Tactics.compute_elim_sig (project g) princ_type in
(* Then we get the number of argument of the function
and compute a fresh name for each of them
*)
- let nb_fun_args = nb_prod (pf_concl g) - 2 in
+ let nb_fun_args = nb_prod (project g) (pf_concl g) - 2 in
let args_names = generate_fresh_id (Id.of_string "x") [] nb_fun_args in
let ids = args_names@(pf_ids_of_hyps g) in
(* and fresh names for res H and the principle (cf bug bug #1174) *)
@@ -685,7 +690,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
(fun decl ->
List.map
(fun id -> id)
- (generate_fresh_id (Id.of_string "y") ids (nb_prod (RelDecl.get_type decl)))
+ (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (RelDecl.get_type decl)))
)
branches
in
@@ -696,7 +701,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
let rewrite_tac j ids : tactic =
let graph_def = graphs.(j) in
let infos =
- try find_Function_infos (fst (destConst funcs.(j)))
+ try find_Function_infos (fst (destConst (project g) funcs.(j)))
with Not_found -> error "No graph found"
in
if infos.is_general
@@ -722,7 +727,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
thin ids
]
else
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))])
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst (project g) f)))])
in
(* The proof of each branche itself *)
let ind_number = ref 0 in
@@ -753,6 +758,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
g
in
let params_names = fst (List.chop princ_infos.nparams args_names) in
+ let open EConstr in
let params = List.map mkVar params_names in
tclTHENSEQ
[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]);
@@ -777,7 +783,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
assert (funs <> []);
assert (graphs <> []);
let funs = Array.of_list funs and graphs = Array.of_list graphs in
- let funs_constr = Array.map mkConstU funs in
+ let map (c, u) = mkConstU (c, EInstance.make u) in
+ let funs_constr = Array.map map funs in
States.with_state_protection_on_exception
(fun () ->
let env = Global.env () in
@@ -792,10 +799,10 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
in
let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in
graphs_constr.(i) <- graph;
- let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in
+ let type_of_lemma = EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in
let _ = Typing.e_type_of (Global.env ()) evd type_of_lemma in
let type_of_lemma = nf_zeta type_of_lemma in
- observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma);
+ observe (str "type_of_lemma := " ++ Printer.pr_leconstr_env (Global.env ()) !evd type_of_lemma);
type_of_lemma,type_info
)
funs_constr
@@ -814,7 +821,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
Array.of_list
(List.map
(fun entry ->
- (fst (fst(Future.force entry.Entries.const_entry_body)), Option.get entry.Entries.const_entry_type )
+ (EConstr.of_constr (fst (fst(Future.force entry.Entries.const_entry_body))), EConstr.of_constr (Option.get entry.Entries.const_entry_type ))
)
(make_scheme evd (Array.map_to_list (fun const -> const,GType []) funs))
)
@@ -845,7 +852,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
(* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *)
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let (lem_cst,_) = destConst lem_cst_constr in
+ let lem_cst_constr = EConstr.of_constr lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with correctness_lemma = Some lem_cst};
)
@@ -859,23 +867,23 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in
graphs_constr.(i) <- graph;
let type_of_lemma =
- Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt
+ EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt
in
let type_of_lemma = nf_zeta type_of_lemma in
- observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma);
+ observe (str "type_of_lemma := " ++ Printer.pr_leconstr type_of_lemma);
type_of_lemma,type_info
)
funs_constr
graphs_constr
in
- let (kn,_) as graph_ind,u = (destInd graphs_constr.(0)) in
+ let (kn,_) as graph_ind,u = (destInd !evd graphs_constr.(0)) in
let mib,mip = Global.lookup_inductive graph_ind in
let sigma, scheme =
(Indrec.build_mutual_induction_scheme (Global.env ()) !evd
(Array.to_list
(Array.mapi
- (fun i _ -> ((kn,i),u(* Univ.Instance.empty *)),true,InType)
+ (fun i _ -> ((kn,i), EInstance.kind !evd u),true,InType)
mib.Declarations.mind_packets
)
)
@@ -905,7 +913,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
let finfo = find_Function_infos (fst f_as_constant) in
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let (lem_cst,_) = destConst lem_cst_constr in
+ let lem_cst_constr = EConstr.of_constr lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with completeness_lemma = Some lem_cst}
)
funs)
@@ -920,10 +929,11 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
if the type of hypothesis has not this form or if we cannot find the completeness lemma then we do nothing
*)
let revert_graph kn post_tac hid g =
+ let sigma = project g in
let typ = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term typ with
- | App(i,args) when isInd i ->
- let ((kn',num) as ind'),u = destInd i in
+ match EConstr.kind sigma typ with
+ | App(i,args) when isInd sigma i ->
+ let ((kn',num) as ind'),u = destInd sigma i in
if MutInd.equal kn kn'
then (* We have generated a graph hypothesis so that we must change it if we can *)
let info =
@@ -971,14 +981,15 @@ let revert_graph kn post_tac hid g =
let functional_inversion kn hid fconst f_correct : tactic =
fun g ->
let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in
+ let sigma = project g in
let type_of_h = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term type_of_h with
- | App(eq,args) when eq_constr eq (make_eq ()) ->
+ match EConstr.kind sigma type_of_h with
+ | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) ->
let pre_tac,f_args,res =
- match kind_of_term args.(1),kind_of_term args.(2) with
- | App(f,f_args),_ when eq_constr f fconst ->
+ match EConstr.kind sigma args.(1),EConstr.kind sigma args.(2) with
+ | App(f,f_args),_ when EConstr.eq_constr sigma f fconst ->
((fun hid -> Proofview.V82.of_tactic (intros_symmetry (Locusops.onHyp hid))),f_args,args.(2))
- |_,App(f,f_args) when eq_constr f fconst ->
+ |_,App(f,f_args) when EConstr.eq_constr sigma f fconst ->
((fun hid -> tclIDTAC),f_args,args.(1))
| _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2)
in
@@ -1023,23 +1034,24 @@ let invfun qhyp f g =
Proofview.V82.of_tactic begin
Tactics.try_intros_until
(fun hid -> Proofview.V82.tactic begin fun g ->
+ let sigma = project g in
let hyp_typ = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term hyp_typ with
- | App(eq,args) when eq_constr eq (make_eq ()) ->
+ match EConstr.kind sigma hyp_typ with
+ | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) ->
begin
- let f1,_ = decompose_app args.(1) in
+ let f1,_ = decompose_app sigma args.(1) in
try
- if not (isConst f1) then failwith "";
- let finfos = find_Function_infos (fst (destConst f1)) in
+ if not (isConst sigma f1) then failwith "";
+ let finfos = find_Function_infos (fst (destConst sigma f1)) in
let f_correct = mkConst(Option.get finfos.correctness_lemma)
and kn = fst finfos.graph_ind
in
functional_inversion kn hid f1 f_correct g
with | Failure "" | Option.IsNone | Not_found ->
try
- let f2,_ = decompose_app args.(2) in
- if not (isConst f2) then failwith "";
- let finfos = find_Function_infos (fst (destConst f2)) in
+ let f2,_ = decompose_app sigma args.(2) in
+ if not (isConst sigma f2) then failwith "";
+ let finfos = find_Function_infos (fst (destConst sigma f2)) in
let f_correct = mkConst(Option.get finfos.correctness_lemma)
and kn = fst finfos.graph_ind
in
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 9c23be68a..f1ca57585 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -32,6 +32,7 @@ module RelDecl = Context.Rel.Declaration
(** {2 Useful operations on constr and glob_constr} *)
+let pop c = Vars.lift (-1) c
let rec popn i c = if i<=0 then c else pop (popn (i-1) c)
(** Substitutions in constr *)
@@ -135,13 +136,14 @@ let prNamedRLDecl s lc =
let showind (id:Id.t) =
let cstrid = Constrintern.global_reference id in
- let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty cstrid in
- let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst ind1) in
+ let (ind1, u),cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty (EConstr.of_constr cstrid) in
+ let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) ind1 in
+ let u = EConstr.Unsafe.to_instance u in
List.iter (fun decl ->
print_string (string_of_name (Context.Rel.Declaration.get_name decl) ^ ":");
prconstr (RelDecl.get_type decl); print_string "\n")
ib1.mind_arity_ctxt;
- Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) ind1);
+ Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) (ind1, u));
Array.iteri
(fun i x -> Printf.printf"type constr %d :" i ; prconstr x)
ib1.mind_user_lc
@@ -776,6 +778,7 @@ let merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
let mkrawcor nme avoid typ =
(* first replace rel 1 by a varname *)
let substindtyp = substitterm 0 (mkRel 1) (mkVar nme) typ in
+ let substindtyp = EConstr.of_constr substindtyp in
Detyping.detype false (Id.Set.elements avoid) (Global.env()) Evd.empty substindtyp in
let lcstr1: glob_constr list =
Array.to_list (Array.map (mkrawcor ind1name avoid) oib1.mind_user_lc) in
@@ -859,6 +862,7 @@ let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) =
match rdecl with
| LocalAssum (nme,t) ->
+ let t = EConstr.of_constr t in
let traw = Detyping.detype false [] (Global.env()) Evd.empty t in
GProd (Loc.ghost,nme,Explicit,traw,t2)
| LocalDef _ -> assert false
@@ -975,23 +979,24 @@ let funify_branches relinfo nfuns branch =
| Rel i -> let reali = i-shift in (reali>=0 && reali<relinfo.nbranches)
| _ -> false in
(* FIXME: *)
- LocalDef (Anonymous,mkProp,mkProp)
+ LocalDef (Anonymous,EConstr.mkProp,EConstr.mkProp)
let relprinctype_to_funprinctype relprinctype nfuns =
- let relinfo = compute_elim_sig relprinctype in
+ let relprinctype = EConstr.of_constr relprinctype in
+ let relinfo = compute_elim_sig Evd.empty (** FIXME*) relprinctype in
assert (not relinfo.farg_in_concl);
assert (relinfo.indarg_in_concl);
(* first remove indarg and indarg_in_concl *)
let relinfo_noindarg = { relinfo with
indarg_in_concl = false; indarg = None;
- concl = remove_last_arg (pop relinfo.concl); } in
+ concl = EConstr.of_constr (remove_last_arg (pop (EConstr.Unsafe.to_constr relinfo.concl))); } in
(* the nfuns last induction arguments are functional ones: remove them *)
let relinfo_argsok = { relinfo_noindarg with
nargs = relinfo_noindarg.nargs - nfuns;
(* args is in reverse order, so remove fst *)
args = remove_n_fst_list nfuns relinfo_noindarg.args;
- concl = popn nfuns relinfo_noindarg.concl
+ concl = EConstr.of_constr (popn nfuns (EConstr.Unsafe.to_constr relinfo_noindarg.concl));
} in
let new_branches =
List.map (funify_branches relinfo_argsok nfuns) relinfo_argsok.branches in
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index e00fa528a..5460d6fb7 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -6,7 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Term
+open EConstr
open Vars
open Namegen
open Environ
@@ -42,17 +45,22 @@ open Indfun_common
open Sigma.Notations
open Context.Rel.Declaration
+let local_assum (na, t) =
+ LocalAssum (na, EConstr.Unsafe.to_constr t)
+
+let local_def (na, b, t) =
+ LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t)
(* Ugly things which should not be here *)
let coq_constant m s =
- Coqlib.coq_constant "RecursiveDefinition" m s
+ EConstr.of_constr (Coqlib.coq_constant "RecursiveDefinition" m s)
let arith_Nat = ["Arith";"PeanoNat";"Nat"]
let arith_Lt = ["Arith";"Lt"]
let coq_init_constant s =
- Coqlib.gen_constant_in_modules "RecursiveDefinition" Coqlib.init_modules s
+ EConstr.of_constr (Coqlib.gen_constant_in_modules "RecursiveDefinition" Coqlib.init_modules s)
let find_reference sl s =
let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in
@@ -76,12 +84,13 @@ let def_of_const t =
)
|_ -> assert false
-let type_of_const t =
- match (kind_of_term t) with
- | Const sp ->
+let type_of_const sigma t =
+ match (EConstr.kind sigma t) with
+ | Const (sp, u) ->
+ let u = EInstance.kind sigma u in
(* FIXME discarding universe constraints *)
- Typeops.type_of_constant_in (Global.env()) sp
- |_ -> assert false
+ Typeops.type_of_constant_in (Global.env()) (sp, u)
+ |_ -> assert false
let constr_of_global x =
fst (Universes.unsafe_constr_of_global x)
@@ -100,9 +109,7 @@ let nf_zeta env =
let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *)
- let clos_norm_flags flgs env sigma t =
- CClosure.norm_val (CClosure.create_clos_infos flgs env) (CClosure.inject (Reductionops.nf_evar sigma t)) in
- clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
+ Reductionops.clos_norm_flags CClosure.betaiotazeta Environ.empty_env Evd.empty
@@ -118,7 +125,7 @@ let pf_get_new_ids idl g =
[]
let compute_renamed_type gls c =
- rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) []
+ rename_bound_vars_as_displayed (project gls) (*no avoid*) [] (*no rels*) []
(pf_unsafe_type_of gls c)
let h'_id = Id.of_string "h'"
let teq_id = Id.of_string "teq"
@@ -149,7 +156,7 @@ let coq_O = function () -> (coq_init_constant "O")
let coq_S = function () -> (coq_init_constant "S")
let lt_n_O = function () -> (coq_constant arith_Nat "nlt_0_r")
let max_ref = function () -> (find_reference ["Recdef"] "max")
-let max_constr = function () -> (constr_of_global (delayed_force max_ref))
+let max_constr = function () -> EConstr.of_constr (constr_of_global (delayed_force max_ref))
let coq_conj = function () -> find_reference Coqlib.logic_module_name "conj"
let f_S t = mkApp(delayed_force coq_S, [|t|]);;
@@ -168,7 +175,8 @@ let simpl_iter clause =
clause
(* Others ugly things ... *)
-let (value_f:constr list -> global_reference -> constr) =
+let (value_f:Constr.constr list -> global_reference -> Constr.constr) =
+ let open Term in
fun al fterm ->
let d0 = Loc.ghost in
let rev_x_id_l =
@@ -201,7 +209,7 @@ let (value_f:constr list -> global_reference -> 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 -> constr list -> global_reference -> global_reference) =
+let (declare_f : Id.t -> logical_kind -> Constr.constr list -> global_reference -> global_reference) =
fun f_id kind input_type fterm_ref ->
declare_fun f_id kind (value_f input_type fterm_ref);;
@@ -303,9 +311,9 @@ let tclUSER_if_not_mes concl_tac is_mes names_to_suppress =
(* [check_not_nested forbidden e] checks that [e] does not contains any variable
of [forbidden]
*)
-let check_not_nested forbidden e =
+let check_not_nested sigma forbidden e =
let rec check_not_nested e =
- match kind_of_term e with
+ match EConstr.kind sigma e with
| Rel _ -> ()
| Var x ->
if Id.List.mem x forbidden
@@ -329,7 +337,7 @@ let check_not_nested forbidden e =
try
check_not_nested e
with UserError(_,p) ->
- user_err ~hdr:"_" (str "on expr : " ++ Printer.pr_lconstr e ++ str " " ++ p)
+ user_err ~hdr:"_" (str "on expr : " ++ Printer.pr_leconstr e ++ str " " ++ p)
(* ['a info] contains the local information for traveling *)
type 'a infos =
@@ -376,15 +384,17 @@ type journey_info =
-let rec add_vars forbidden e =
- match kind_of_term e with
+let add_vars sigma forbidden e =
+ let rec aux forbidden e =
+ match EConstr.kind sigma e with
| Var x -> x::forbidden
- | _ -> Term.fold_constr add_vars forbidden e
-
+ | _ -> EConstr.fold sigma aux forbidden e
+ in
+ aux forbidden e
let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
fun g ->
- let rev_context,b = decompose_lam_n nb_lam e in
+ let rev_context,b = decompose_lam_n (project g) nb_lam e in
let ids = List.fold_left (fun acc (na,_) ->
let pre_id =
match na with
@@ -406,17 +416,17 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
(fun g' ->
let ty_teq = pf_unsafe_type_of g' (mkVar heq) in
let teq_lhs,teq_rhs =
- let _,args = try destApp ty_teq with DestKO -> assert false in
+ let _,args = try destApp (project g') ty_teq with DestKO -> assert false in
args.(1),args.(2)
in
- let new_b' = Termops.replace_term teq_lhs teq_rhs new_b in
+ let new_b' = Termops.replace_term (project g') teq_lhs teq_rhs new_b in
let new_infos = {
infos with
info = new_b';
eqs = heq::infos.eqs;
forbidden_ids =
if forbid_new_ids
- then add_vars infos.forbidden_ids new_b'
+ then add_vars (project g') infos.forbidden_ids new_b'
else infos.forbidden_ids
} in
finalize_tac new_infos g'
@@ -425,8 +435,9 @@ let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
)
] g
-let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
- match kind_of_term expr_info.info with
+let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
+ let sigma = project g in
+ match EConstr.kind sigma expr_info.info with
| CoFix _ | Fix _ -> error "Function cannot treat local fixpoint or cofixpoint"
| Proj _ -> error "Function cannot treat projections"
| LetIn(na,b,t,e) ->
@@ -435,24 +446,24 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
jinfo.letiN (na,b,t,e) expr_info continuation_tac
in
travel jinfo new_continuation_tac
- {expr_info with info = b; is_final=false}
+ {expr_info with info = b; is_final=false} g
end
| Rel _ -> anomaly (Pp.str "Free var in goal conclusion !")
| Prod _ ->
begin
try
- check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
- jinfo.otherS () expr_info continuation_tac expr_info
+ check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ jinfo.otherS () expr_info continuation_tac expr_info g
with e when CErrors.noncritical e ->
- user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
+ user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
end
| Lambda(n,t,b) ->
begin
try
- check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
- jinfo.otherS () expr_info continuation_tac expr_info
+ check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ jinfo.otherS () expr_info continuation_tac expr_info g
with e when CErrors.noncritical e ->
- user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
+ user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
end
| Case(ci,t,a,l) ->
begin
@@ -463,15 +474,15 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
travel
jinfo continuation_tac_a
{expr_info with info = a; is_main_branch = false;
- is_final = false}
+ is_final = false} g
end
| App _ ->
- let f,args = decompose_app expr_info.info in
- if eq_constr f (expr_info.f_constr)
- then jinfo.app_reC (f,args) expr_info continuation_tac expr_info
+ let f,args = decompose_app sigma expr_info.info in
+ if EConstr.eq_constr sigma f (expr_info.f_constr)
+ then jinfo.app_reC (f,args) expr_info continuation_tac expr_info g
else
begin
- match kind_of_term f with
+ match EConstr.kind sigma f with
| App _ -> assert false (* f is coming from a decompose_app *)
| Const _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _
| Sort _ | Prod _ | Var _ ->
@@ -479,15 +490,15 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
let new_continuation_tac =
jinfo.apP (f,args) expr_info continuation_tac in
travel_args jinfo
- expr_info.is_main_branch new_continuation_tac new_infos
- | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
- | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_lconstr expr_info.info)
+ expr_info.is_main_branch new_continuation_tac new_infos g
+ | Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
+ | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr expr_info.info)
end
- | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t}
+ | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g
| Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ ->
let new_continuation_tac =
jinfo.otherS () expr_info continuation_tac in
- new_continuation_tac expr_info
+ new_continuation_tac expr_info g
and travel_args jinfo is_final continuation_tac infos =
let (f_args',args) = infos.info in
match args with
@@ -504,27 +515,28 @@ and travel_args jinfo is_final continuation_tac infos =
{infos with info=arg;is_final=false}
and travel jinfo continuation_tac expr_info =
observe_tac
- (str jinfo.message ++ Printer.pr_lconstr expr_info.info)
+ (str jinfo.message ++ Printer.pr_leconstr expr_info.info)
(travel_aux jinfo continuation_tac expr_info)
(* Termination proof *)
let rec prove_lt hyple g =
+ let sigma = project g in
begin
try
- let (varx,varz) = match decompose_app (pf_concl g) with
- | _, x::z::_ when isVar x && isVar z -> x, z
+ let (varx,varz) = match decompose_app sigma (pf_concl g) with
+ | _, x::z::_ when isVar sigma x && isVar sigma z -> x, z
| _ -> assert false
in
let h =
List.find (fun id ->
- match decompose_app (pf_unsafe_type_of g (mkVar id)) with
- | _, t::_ -> eq_constr t varx
+ match decompose_app sigma (pf_unsafe_type_of g (mkVar id)) with
+ | _, t::_ -> EConstr.eq_constr sigma t varx
| _ -> false
) hyple
in
let y =
- List.hd (List.tl (snd (decompose_app (pf_unsafe_type_of g (mkVar h))))) in
+ List.hd (List.tl (snd (decompose_app sigma (pf_unsafe_type_of g (mkVar h))))) in
observe_tclTHENLIST (str "prove_lt1")[
Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|])));
observe_tac (str "prove_lt") (prove_lt hyple)
@@ -640,12 +652,13 @@ let terminate_others _ expr_info continuation_tac infos =
]
else continuation_tac infos
-let terminate_letin (na,b,t,e) expr_info continuation_tac info =
+let terminate_letin (na,b,t,e) expr_info continuation_tac info g =
+ let sigma = project g in
let new_e = subst1 info.info e in
let new_forbidden =
let forbid =
try
- check_not_nested (expr_info.f_id::expr_info.forbidden_ids) b;
+ check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) b;
true
with e when CErrors.noncritical e -> false
in
@@ -656,7 +669,7 @@ let terminate_letin (na,b,t,e) expr_info continuation_tac info =
| Name id -> id::info.forbidden_ids
else info.forbidden_ids
in
- continuation_tac {info with info = new_e; forbidden_ids = new_forbidden}
+ continuation_tac {info with info = new_e; forbidden_ids = new_forbidden} g
let pf_type c tac gl =
let evars, ty = Typing.type_of (pf_env gl) (project gl) c in
@@ -683,7 +696,7 @@ let mkDestructEq :
(fun decl ->
let open Context.Named.Declaration in
let id = get_id decl in
- if Id.List.mem id not_on_hyp || not (Termops.occur_term expr (get_type decl))
+ if Id.List.mem id not_on_hyp || not (Termops.occur_term (project g) expr (get_type decl))
then None else Some id) hyps in
let to_revert_constr = List.rev_map mkVar to_revert in
let type_of_expr = pf_unsafe_type_of g expr in
@@ -695,16 +708,18 @@ let mkDestructEq :
(fun g2 ->
let changefun patvars = { run = fun sigma ->
let redfun = pattern_occs [Locus.AllOccurrencesBut [1], expr] in
- redfun.Reductionops.e_redfun (pf_env g2) sigma (pf_concl g2)
+ let Sigma (c, sigma, p) = redfun.Reductionops.e_redfun (pf_env g2) sigma (pf_concl g2) in
+ Sigma (c, sigma, p)
} in
Proofview.V82.of_tactic (change_in_concl None changefun) g2);
Proofview.V82.of_tactic (simplest_case expr)]), to_revert
let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
+ let sigma = project g in
let f_is_present =
try
- check_not_nested (expr_info.f_id::expr_info.forbidden_ids) a;
+ check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids) a;
false
with e when CErrors.noncritical e ->
true
@@ -718,7 +733,7 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
let destruct_tac,rev_to_thin_intro =
mkDestructEq [expr_info.rec_arg_id] a' g in
let to_thin_intro = List.rev rev_to_thin_intro in
- observe_tac (str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_lconstr a')
+ observe_tac (str "treating cases (" ++ int (Array.length l) ++ str")" ++ spc () ++ Printer.pr_leconstr a')
(try
(tclTHENS
destruct_tac
@@ -727,16 +742,17 @@ let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
with
| UserError(Some "Refiner.thensn_tac3",_)
| UserError(Some "Refiner.tclFAIL_s",_) ->
- (observe_tac (str "is computable " ++ Printer.pr_lconstr new_info.info) (next_step continuation_tac {new_info with info = nf_betaiotazeta new_info.info} )
+ (observe_tac (str "is computable " ++ Printer.pr_leconstr new_info.info) (next_step continuation_tac {new_info with info = nf_betaiotazeta new_info.info} )
))
g
-let terminate_app_rec (f,args) expr_info continuation_tac _ =
- List.iter (check_not_nested (expr_info.f_id::expr_info.forbidden_ids))
+let terminate_app_rec (f,args) expr_info continuation_tac _ g =
+ let sigma = project g in
+ List.iter (check_not_nested sigma (expr_info.f_id::expr_info.forbidden_ids))
args;
begin
try
- let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in
+ let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in
let new_infos = {expr_info with info = v} in
observe_tclTHENLIST (str "terminate_app_rec")[
continuation_tac new_infos;
@@ -750,7 +766,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ =
]
else
tclIDTAC
- ]
+ ] g
with Not_found ->
observe_tac (str "terminate_app_rec not found") (tclTHENS
(Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args))))
@@ -807,7 +823,7 @@ let terminate_app_rec (f,args) expr_info continuation_tac _ =
);
]
])
- ])
+ ]) g
end
let terminate_info =
@@ -829,8 +845,9 @@ let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos =
observe_tac (str "equation case") (terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos)
let rec prove_le g =
+ let sigma = project g in
let x,z =
- let _,args = decompose_app (pf_concl g) in
+ let _,args = decompose_app sigma (pf_concl g) in
(List.hd args,List.hd (List.tl args))
in
tclFIRST[
@@ -840,11 +857,11 @@ let rec prove_le g =
try
let matching_fun =
pf_is_matching g
- (Pattern.PApp(Pattern.PRef (reference_of_constr (le ())),[|Pattern.PVar (destVar x);Pattern.PMeta None|])) in
+ (Pattern.PApp(Pattern.PRef (reference_of_constr (EConstr.Unsafe.to_constr (le ()))),[|Pattern.PVar (destVar sigma x);Pattern.PMeta None|])) in
let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g)
in
let y =
- let _,args = decompose_app t in
+ let _,args = decompose_app sigma t in
List.hd (List.tl args)
in
observe_tclTHENLIST (str "prove_le")[
@@ -862,11 +879,12 @@ let rec make_rewrite_list expr_info max = function
observe_tac (str "make_rewrite_list") (tclTHENS
(observe_tac (str "rewrite heq on " ++ pr_id p ) (
(fun g ->
+ let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
let k,def =
- let k_na,_,t = destProd t_eq in
- let _,_,t = destProd t in
- let def_na,_,_ = destProd t in
+ let k_na,_,t = destProd sigma t_eq in
+ let _,_,t = destProd sigma t in
+ let def_na,_,_ = destProd sigma t in
Nameops.out_name k_na,Nameops.out_name def_na
in
Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
@@ -874,7 +892,7 @@ let rec make_rewrite_list expr_info max = function
(mkVar hp,
ExplicitBindings[Loc.ghost,NamedHyp def,
expr_info.f_constr;Loc.ghost,NamedHyp k,
- (f_S max)]) false) g) )
+ f_S max]) false) g) )
)
[make_rewrite_list expr_info max l;
observe_tclTHENLIST (str "make_rewrite_list")[ (* x < S max proof *)
@@ -888,11 +906,12 @@ let make_rewrite expr_info l hp max =
(observe_tac (str "make_rewrite") (make_rewrite_list expr_info max l))
(observe_tac (str "make_rewrite") (tclTHENS
(fun g ->
+ let sigma = project g in
let t_eq = compute_renamed_type g (mkVar hp) in
let k,def =
- let k_na,_,t = destProd t_eq in
- let _,_,t = destProd t in
- let def_na,_,_ = destProd t in
+ let k_na,_,t = destProd sigma t_eq in
+ let _,_,t = destProd sigma t in
+ let def_na,_,_ = destProd sigma t in
Nameops.out_name k_na,Nameops.out_name def_na
in
observe_tac (str "general_rewrite_bindings")
@@ -901,7 +920,7 @@ let make_rewrite expr_info l hp max =
(mkVar hp,
ExplicitBindings[Loc.ghost,NamedHyp def,
expr_info.f_constr;Loc.ghost,NamedHyp k,
- (f_S (f_S max))]) false)) g)
+ f_S (f_S max)]) false)) g)
[observe_tac(str "make_rewrite finalize") (
(* tclORELSE( h_reflexivity) *)
(observe_tclTHENLIST (str "make_rewrite")[
@@ -918,7 +937,7 @@ let make_rewrite expr_info l hp max =
]))
;
observe_tclTHENLIST (str "make_rewrite1")[ (* x < S (S max) proof *)
- Proofview.V82.of_tactic (apply (delayed_force le_lt_SS));
+ Proofview.V82.of_tactic (apply (EConstr.of_constr (delayed_force le_lt_SS)));
observe_tac (str "prove_le (3)") prove_le
]
])
@@ -976,23 +995,24 @@ let rec intros_values_eq expr_info acc =
let equation_others _ expr_info continuation_tac infos =
if expr_info.is_final && expr_info.is_main_branch
then
- observe_tac (str "equation_others (cont_tac +intros) " ++ Printer.pr_lconstr expr_info.info)
+ observe_tac (str "equation_others (cont_tac +intros) " ++ Printer.pr_leconstr expr_info.info)
(tclTHEN
(continuation_tac infos)
- (observe_tac (str "intros_values_eq equation_others " ++ Printer.pr_lconstr expr_info.info) (intros_values_eq expr_info [])))
- else observe_tac (str "equation_others (cont_tac) " ++ Printer.pr_lconstr expr_info.info) (continuation_tac infos)
+ (observe_tac (str "intros_values_eq equation_others " ++ Printer.pr_leconstr expr_info.info) (intros_values_eq expr_info [])))
+ else observe_tac (str "equation_others (cont_tac) " ++ Printer.pr_leconstr expr_info.info) (continuation_tac infos)
let equation_app f_and_args expr_info continuation_tac infos =
if expr_info.is_final && expr_info.is_main_branch
then ((observe_tac (str "intros_values_eq equation_app") (intros_values_eq expr_info [])))
else continuation_tac infos
-let equation_app_rec (f,args) expr_info continuation_tac info =
+let equation_app_rec (f,args) expr_info continuation_tac info g =
+ let sigma = project g in
begin
try
- let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in
+ let v = List.assoc_f (List.equal (EConstr.eq_constr sigma)) args expr_info.args_assoc in
let new_infos = {expr_info with info = v} in
- observe_tac (str "app_rec found") (continuation_tac new_infos)
+ observe_tac (str "app_rec found") (continuation_tac new_infos) g
with Not_found ->
if expr_info.is_final && expr_info.is_main_branch
then
@@ -1000,12 +1020,12 @@ let equation_app_rec (f,args) expr_info continuation_tac info =
[ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc};
observe_tac (str "app_rec intros_values_eq") (intros_values_eq expr_info [])
- ]
+ ] g
else
observe_tclTHENLIST (str "equation_app_rec1")[
Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
observe_tac (str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc})
- ]
+ ] g
end
let equation_info =
@@ -1024,6 +1044,8 @@ let prove_eq = travel equation_info
(* [compute_terminate_type] computes the type of the Definition f_terminate from the type of f_F
*)
let compute_terminate_type nb_args func =
+ let open Term 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 =
@@ -1036,6 +1058,7 @@ let compute_terminate_type nb_args func =
)
in
let right = mkRel 5 in
+ let delayed_force c = EConstr.Unsafe.to_constr (delayed_force c) in
let equality = mkApp(delayed_force eq, [|lift 5 b; left; right|]) in
let result = (mkProd ((Name def_id) , lift 4 a_arrow_b, equality)) in
let cond = mkApp(delayed_force lt, [|(mkRel 2); (mkRel 1)|]) in
@@ -1048,7 +1071,7 @@ let compute_terminate_type nb_args func =
delayed_force nat,
(mkProd (Name k_id, delayed_force nat,
mkArrow cond result))))|])in
- let value = mkApp(constr_of_global (delayed_force coq_sig_ref),
+ let value = mkApp(constr_of_global (Util.delayed_force coq_sig_ref),
[|b;
(mkLambda (Name v_id, b, nb_iter))|]) in
compose_prod rev_args value
@@ -1132,25 +1155,27 @@ let termination_proof_header is_mes input_type ids args_id relation
-let rec instantiate_lambda t l =
+let rec instantiate_lambda sigma t l =
match l with
| [] -> t
| a::l ->
- let (_, _, body) = destLambda t in
- instantiate_lambda (subst1 a body) l
+ let (_, _, body) = destLambda sigma t in
+ instantiate_lambda sigma (subst1 a body) l
let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_arg_num : tactic =
begin
fun g ->
+ let sigma = project g in
let ids = Termops.ids_of_named_context (pf_hyps g) in
let func_body = (def_of_const (constr_of_global func)) in
- let (f_name, _, body1) = destLambda func_body in
+ let func_body = EConstr.of_constr func_body in
+ let (f_name, _, body1) = destLambda sigma func_body in
let f_id =
match f_name with
| Name f_id -> next_ident_away_in_goal f_id ids
| Anonymous -> anomaly (Pp.str "Anonymous function")
in
- let n_names_types,_ = decompose_lam_n nb_args body1 in
+ let n_names_types,_ = decompose_lam_n sigma nb_args body1 in
let n_ids,ids =
List.fold_left
(fun (n_ids,ids) (n_name,_) ->
@@ -1164,7 +1189,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
n_names_types
in
let rec_arg_id = List.nth n_ids (rec_arg_num - 1) in
- let expr = instantiate_lambda func_body (mkVar f_id::(List.map mkVar n_ids)) in
+ let expr = instantiate_lambda sigma func_body (mkVar f_id::(List.map mkVar n_ids)) in
termination_proof_header
is_mes
input_type
@@ -1206,17 +1231,17 @@ let get_current_subgoals_types () =
let { Evd.it=sgs ; sigma=sigma } = Proof.V82.subgoals p in
sigma, List.map (Goal.V82.abstract_type sigma) sgs
-let build_and_l l =
+let build_and_l sigma l =
let and_constr = Coqlib.build_coq_and () in
let conj_constr = coq_conj () in
let mk_and p1 p2 =
- Term.mkApp(and_constr,[|p1;p2|]) in
+ mkApp(EConstr.of_constr and_constr,[|p1;p2|]) in
let rec is_well_founded t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod(_,_,t') -> is_well_founded t'
| App(_,_) ->
- let (f,_) = decompose_app t in
- eq_constr f (well_founded ())
+ let (f,_) = decompose_app sigma t in
+ EConstr.eq_constr sigma f (well_founded ())
| _ ->
false
in
@@ -1233,7 +1258,7 @@ let build_and_l l =
let c,tac,nb = f pl in
mk_and p1 c,
tclTHENS
- (Proofview.V82.of_tactic (apply (constr_of_global conj_constr)))
+ (Proofview.V82.of_tactic (apply (EConstr.of_constr (constr_of_global conj_constr))))
[tclIDTAC;
tac
],nb+1
@@ -1247,16 +1272,16 @@ let is_rec_res id =
String.equal (String.sub id_name 0 (String.length rec_res_name)) rec_res_name
with Invalid_argument _ -> false
-let clear_goals =
+let clear_goals sigma =
let rec clear_goal t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod(Name id as na,t',b) ->
let b' = clear_goal b in
- if noccurn 1 b' && (is_rec_res id)
- then Termops.pop b'
+ if noccurn sigma 1 b' && (is_rec_res id)
+ then Vars.lift (-1) b'
else if b' == b then t
else mkProd(na,t',b')
- | _ -> Term.map_constr clear_goal t
+ | _ -> EConstr.map sigma clear_goal t
in
List.map clear_goal
@@ -1264,9 +1289,9 @@ let clear_goals =
let build_new_goal_type () =
let sigma, sub_gls_types = get_current_subgoals_types () in
(* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
- let sub_gls_types = clear_goals sub_gls_types in
+ let sub_gls_types = clear_goals sigma sub_gls_types in
(* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
- let res = build_and_l sub_gls_types in
+ let res = build_and_l sigma sub_gls_types in
sigma, res
let is_opaque_constant c =
@@ -1287,7 +1312,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
anomaly (Pp.str "open_new_goal with an unamed theorem")
in
let na = next_global_ident_away name [] in
- if Termops.occur_existential gls_type then
+ if Termops.occur_existential sigma gls_type then
CErrors.error "\"abstract\" cannot handle existentials";
let hook _ _ =
let opacity =
@@ -1298,7 +1323,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
| _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant")
in
let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in
- ref_ := Some lemma ;
+ ref_ := Some (EConstr.Unsafe.to_constr lemma);
let lid = ref [] in
let h_num = ref (-1) in
let env = Global.env () in
@@ -1324,8 +1349,9 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
);
] gls)
(fun g ->
- match kind_of_term (pf_concl g) with
- | App(f,_) when eq_constr f (well_founded ()) ->
+ let sigma = project g in
+ match EConstr.kind sigma (pf_concl g) with
+ | App(f,_) when EConstr.eq_constr sigma f (well_founded ()) ->
Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g
| _ ->
incr h_num;
@@ -1368,7 +1394,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
(fun c ->
Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST
[intros;
- Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*);
+ Simple.apply (EConstr.of_constr (fst (interp_constr (Global.env()) Evd.empty c))) (*FIXME*);
Tacticals.New.tclCOMPLETE Auto.default_auto
])
)
@@ -1385,7 +1411,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
let com_terminate
tcc_lemma_name
- tcc_lemma_ref
+ (tcc_lemma_ref : Constr.t option ref)
is_mes
fonctional_ref
input_type
@@ -1398,7 +1424,7 @@ let com_terminate
let (evmap, env) = Lemmas.get_current_context() in
Lemmas.start_proof thm_name
(Global, false (* FIXME *), Proof Lemma) ~sign:(Environ.named_context_val env)
- ctx (compute_terminate_type nb_args fonctional_ref) hook;
+ ctx (EConstr.of_constr (compute_terminate_type nb_args fonctional_ref)) hook;
ignore (by (Proofview.V82.tactic (observe_tac (str "starting_tac") tac_start)));
ignore (by (Proofview.V82.tactic (observe_tac (str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref
@@ -1422,9 +1448,11 @@ let com_terminate
let start_equation (f:global_reference) (term_f:global_reference)
(cont_tactic:Id.t list -> tactic) g =
+ let sigma = project g in
let ids = pf_ids_of_hyps g in
let terminate_constr = constr_of_global term_f in
- let nargs = nb_prod (type_of_const terminate_constr) in
+ let terminate_constr = EConstr.of_constr terminate_constr in
+ let nargs = nb_prod (project g) (EConstr.of_constr (type_of_const sigma terminate_constr)) in
let x = n_x_id ids nargs in
observe_tac (str "start_equation") (observe_tclTHENLIST (str "start_equation") [
h_intros x;
@@ -1436,8 +1464,9 @@ let start_equation (f:global_reference) (term_f:global_reference)
let (com_eqn : int -> Id.t ->
global_reference -> global_reference -> global_reference
- -> constr -> unit) =
+ -> Constr.constr -> unit) =
fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type ->
+ let open CVars in
let opacity =
match terminate_ref with
| ConstRef c -> is_opaque_constant c
@@ -1450,20 +1479,20 @@ let (com_eqn : int -> Id.t ->
(Lemmas.start_proof eq_name (Global, false, Proof Lemma)
~sign:(Environ.named_context_val env)
evmap
- equation_lemma_type
+ (EConstr.of_constr equation_lemma_type)
(Lemmas.mk_hook (fun _ _ -> ()));
ignore (by
(Proofview.V82.tactic (start_equation f_ref terminate_ref
(fun x ->
prove_eq (fun _ -> tclIDTAC)
{nb_arg=nb_arg;
- f_terminate = constr_of_global terminate_ref;
- f_constr = f_constr;
+ f_terminate = EConstr.of_constr (constr_of_global terminate_ref);
+ f_constr = EConstr.of_constr f_constr;
concl_tac = tclIDTAC;
func=functional_ref;
- info=(instantiate_lambda
- (def_of_const (constr_of_global functional_ref))
- (f_constr::List.map mkVar x)
+ info=(instantiate_lambda Evd.empty
+ (EConstr.of_constr (def_of_const (constr_of_global functional_ref)))
+ (EConstr.of_constr f_constr::List.map mkVar x)
);
is_main_branch = true;
is_final = true;
@@ -1489,19 +1518,25 @@ 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 CVars in
let env = Global.env() in
let evd = ref (Evd.from_env env) in
let function_type = interp_type_evars env evd type_of_f in
+ let function_type = EConstr.Unsafe.to_constr function_type in
let env = push_named (Context.Named.Declaration.LocalAssum (function_name,function_type)) env in
(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *)
let ty = interp_type_evars env evd ~impls:rec_impls eq in
+ let ty = EConstr.Unsafe.to_constr ty in
let evm, nf = Evarutil.nf_evars_and_universes !evd in
- let equation_lemma_type = nf_betaiotazeta (nf ty) in
+ let equation_lemma_type = nf_betaiotazeta (EConstr.of_constr (nf ty)) in
let function_type = nf function_type in
+ let equation_lemma_type = EConstr.Unsafe.to_constr equation_lemma_type in
(* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *)
let res_vars,eq' = decompose_prod equation_lemma_type in
let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> LocalAssum (x,y)) res_vars) env in
- let eq' = nf_zeta env_eq' eq' in
+ let eq' = nf_zeta env_eq' (EConstr.of_constr eq') in
+ let eq' = EConstr.Unsafe.to_constr eq' in
let res =
(* 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)); *)
@@ -1554,7 +1589,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
and functional_ref = destConst (constr_of_global functional_ref)
and eq_ref = destConst (constr_of_global eq_ref) in
generate_induction_principle f_ref tcc_lemma_constr
- functional_ref eq_ref rec_arg_num rec_arg_type (nb_prod res) relation;
+ functional_ref eq_ref rec_arg_num (EConstr.of_constr rec_arg_type) (nb_prod evm (EConstr.of_constr res)) (EConstr.of_constr relation);
if Flags.is_verbose ()
then msgnl (h 1 (Ppconstr.pr_id function_name ++
spc () ++ str"is defined" )++ fnl () ++
@@ -1567,8 +1602,8 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
tcc_lemma_name
tcc_lemma_constr
is_mes functional_ref
- rec_arg_type
- relation rec_arg_num
+ (EConstr.of_constr rec_arg_type)
+ (EConstr.of_constr relation) rec_arg_num
term_id
using_lemmas
(List.length res_vars)
diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli
index f60eedbe6..9c1081b9d 100644
--- a/plugins/funind/recdef.mli
+++ b/plugins/funind/recdef.mli
@@ -15,6 +15,6 @@ bool ->
int -> Constrexpr.constr_expr -> (Term.pconstant ->
Term.constr option ref ->
Term.pconstant ->
- Term.pconstant -> int -> Term.types -> int -> Term.constr -> 'a) -> Constrexpr.constr_expr list -> unit
+ Term.pconstant -> int -> EConstr.types -> int -> EConstr.constr -> 'a) -> Constrexpr.constr_expr list -> unit
diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml
index c5b26e6d5..5d3f6df03 100644
--- a/plugins/ltac/evar_tactics.ml
+++ b/plugins/ltac/evar_tactics.ml
@@ -7,6 +7,9 @@
(************************************************************************)
open Util
+open Names
+open Term
+open EConstr
open CErrors
open Evar_refiner
open Tacmach
@@ -35,25 +38,32 @@ let instantiate_evar evk (ist,rawc) sigma =
let sigma' = w_refine (evk,evi) (lvar ,rawc) sigma in
tclEVARS sigma'
+let evar_list sigma c =
+ let rec evrec acc c =
+ match EConstr.kind sigma c with
+ | Evar (evk, _ as ev) -> ev :: acc
+ | _ -> EConstr.fold sigma evrec acc c in
+ evrec [] c
+
let instantiate_tac n c ido =
Proofview.V82.tactic begin fun gl ->
let sigma = gl.sigma in
let evl =
match ido with
- ConclLocation () -> evar_list (pf_concl gl)
+ ConclLocation () -> evar_list sigma (pf_concl gl)
| HypLocation (id,hloc) ->
let decl = Environ.lookup_named_val id (Goal.V82.hyps sigma (sig_it gl)) in
match hloc with
InHyp ->
(match decl with
- | LocalAssum (_,typ) -> evar_list typ
+ | LocalAssum (_,typ) -> evar_list sigma (EConstr.of_constr typ)
| _ -> error
"Please be more specific: in type or value?")
| InHypTypeOnly ->
- evar_list (NamedDecl.get_type decl)
+ evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl))
| InHypValueOnly ->
(match decl with
- | LocalDef (_,body,_) -> evar_list body
+ | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body)
| _ -> error "Not a defined hypothesis.") in
if List.length evl < n then
error "Not enough uninstantiated existential variables.";
@@ -78,16 +88,32 @@ let let_evar name typ =
let env = Proofview.Goal.env gl in
let sigma = ref sigma in
let _ = Typing.e_sort_of env sigma typ in
- let sigma = Sigma.Unsafe.of_evar_map !sigma in
+ let sigma = !sigma in
let id = match name with
| Names.Anonymous ->
- let id = Namegen.id_of_name_using_hdchar env typ name in
+ let id = Namegen.id_of_name_using_hdchar env sigma typ name in
Namegen.next_ident_away_in_goal id (Termops.ids_of_named_context (Environ.named_context env))
| Names.Name id -> id
in
+ let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (evar, sigma, p) = Evarutil.new_evar env sigma ~src ~naming:(Misctypes.IntroFresh id) typ in
let tac =
(Tactics.letin_tac None (Names.Name id) evar None Locusops.nowhere)
in
Sigma (tac, sigma, p)
end }
+
+let hget_evar n =
+ let open EConstr in
+ Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
+ let concl = Proofview.Goal.concl gl in
+ let evl = evar_list sigma concl in
+ if List.length evl < n then
+ error "Not enough uninstantiated existential variables.";
+ if n <= 0 then error "Incorrect existential variable index.";
+ let ev = List.nth evl (n-1) in
+ let ev_type = EConstr.existential_type sigma ev in
+ Tactics.change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl))
+ end }
+
diff --git a/plugins/ltac/evar_tactics.mli b/plugins/ltac/evar_tactics.mli
index e67540c05..cfe747665 100644
--- a/plugins/ltac/evar_tactics.mli
+++ b/plugins/ltac/evar_tactics.mli
@@ -16,4 +16,6 @@ val instantiate_tac : int -> Tacinterp.interp_sign * Glob_term.glob_constr ->
val instantiate_tac_by_name : Id.t ->
Tacinterp.interp_sign * Glob_term.glob_constr -> unit Proofview.tactic
-val let_evar : Name.t -> Term.types -> unit Proofview.tactic
+val let_evar : Name.t -> EConstr.types -> unit Proofview.tactic
+
+val hget_evar : int -> unit Proofview.tactic
diff --git a/plugins/ltac/extraargs.mli b/plugins/ltac/extraargs.mli
index b12187e18..7d4bccfad 100644
--- a/plugins/ltac/extraargs.mli
+++ b/plugins/ltac/extraargs.mli
@@ -38,12 +38,12 @@ val wit_lglob :
val wit_lconstr :
(constr_expr,
Tacexpr.glob_constr_and_expr,
- Constr.t) Genarg.genarg_type
+ EConstr.t) Genarg.genarg_type
val wit_casted_constr :
(constr_expr,
Tacexpr.glob_constr_and_expr,
- Constr.t) Genarg.genarg_type
+ EConstr.t) Genarg.genarg_type
val glob : constr_expr Pcoq.Gram.entry
val lglob : constr_expr Pcoq.Gram.entry
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 7a9fc6657..38fdfb759 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -118,7 +118,7 @@ END
let discrHyp id =
Proofview.tclEVARMAP >>= fun sigma ->
- discr_main { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma }
+ discr_main { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma }
let injection_main with_evars c =
elimOnConstrWithHoles (injClause None) with_evars c
@@ -150,7 +150,7 @@ END
let injHyp id =
Proofview.tclEVARMAP >>= fun sigma ->
- injection_main false { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma }
+ injection_main false { delayed = fun env sigma -> Sigma.here (EConstr.mkVar id, NoBindings) sigma }
TACTIC EXTEND dependent_rewrite
| [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ]
@@ -290,6 +290,7 @@ END
(* Hint Resolve *)
open Term
+open EConstr
open Vars
open Coqlib
@@ -298,22 +299,25 @@ let project_hint pri l2r r =
let env = Global.env() in
let sigma = Evd.from_env env in
let sigma, c = Evd.fresh_global env sigma gr in
+ let c = EConstr.of_constr c in
let t = Retyping.get_type_of env sigma c in
let t =
Tacred.reduce_to_quantified_ref env sigma (Lazy.force coq_iff_ref) t in
- let sign,ccl = decompose_prod_assum t in
- let (a,b) = match snd (decompose_app ccl) with
+ let sign,ccl = decompose_prod_assum sigma t in
+ let (a,b) = match snd (decompose_app sigma ccl) with
| [a;b] -> (a,b)
| _ -> assert false in
let p =
if l2r then build_coq_iff_left_proj () else build_coq_iff_right_proj () in
- let c = Reductionops.whd_beta Evd.empty (mkApp (c, Context.Rel.to_extended_vect 0 sign)) in
+ let p = EConstr.of_constr p in
+ let c = Reductionops.whd_beta sigma (mkApp (c, Context.Rel.to_extended_vect mkRel 0 sign)) in
let c = it_mkLambda_or_LetIn
(mkApp (p,[|mkArrow a (lift 1 b);mkArrow b (lift 1 a);c|])) sign in
let id =
Nameops.add_suffix (Nametab.basename_of_global gr) ("_proj_" ^ (if l2r then "l2r" else "r2l"))
in
let ctx = Evd.universe_context_set sigma in
+ let c = EConstr.to_constr sigma c in
let c = Declare.declare_definition ~internal:Declare.InternalTacticRequest id (c,ctx) in
let info = {Vernacexpr.hint_priority = pri; hint_pattern = None} in
(info,false,true,Hints.PathAny, Hints.IsGlobRef (Globnames.ConstRef c))
@@ -341,6 +345,9 @@ END
(**********************************************************************)
(* Refine *)
+open EConstr
+open Vars
+
let constr_flags () = {
Pretyping.use_typeclasses = true;
Pretyping.solve_unification_constraints = true;
@@ -349,14 +356,17 @@ let constr_flags () = {
Pretyping.expand_evars = true }
let refine_tac ist simple with_classes c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
let flags =
{ constr_flags () with Pretyping.use_typeclasses = with_classes } in
let expected_type = Pretyping.OfType concl in
let c = Pretyping.type_uconstr ~flags ~expected_type ist c in
- let update = { run = fun sigma -> c.delayed env sigma } in
+ let update = { run = fun sigma ->
+ let Sigma (c, sigma, p) = c.delayed env sigma in
+ Sigma (c, sigma, p)
+ } in
let refine = Refine.refine ~unsafe:true update in
if simple then refine
else refine <*>
@@ -412,8 +422,6 @@ VERNAC COMMAND EXTEND DeriveInversionClear
-> [ add_inversion_lemma_exn na c GProp false inv_clear_tac ]
END
-open Term
-
VERNAC COMMAND EXTEND DeriveInversion
| [ "Derive" "Inversion" ident(na) "with" constr(c) "Sort" sort(s) ]
=> [ seff na ]
@@ -492,6 +500,7 @@ let transitivity_left_table = Summary.ref [] ~name:"transitivity-steps-l"
let step left x tac =
let l =
List.map (fun lem ->
+ let lem = EConstr.of_constr lem in
Tacticals.New.tclTHENLAST
(apply_with_bindings (lem, ImplicitBindings [x]))
tac)
@@ -509,7 +518,7 @@ let cache_transitivity_lemma (_,(left,lem)) =
let subst_transitivity_lemma (subst,(b,ref)) = (b,subst_mps subst ref)
-let inTransitivity : bool * constr -> obj =
+let inTransitivity : bool * Constr.constr -> 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);
@@ -656,7 +665,7 @@ let subst_hole_with_term occ tc t =
open Tacmach
let hResolve id c occ t =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let sigma = Sigma.to_evar_map sigma in
let env = Termops.clear_named_body id (Proofview.Goal.env gl) in
@@ -674,6 +683,7 @@ let hResolve id c occ t =
resolve_hole (subst_hole_with_term (fst (Loc.unloc loc)) c_raw t_hole)
in
let t_constr,ctx = resolve_hole (subst_var_with_hole occ id t_raw) in
+ let t_constr = EConstr.of_constr t_constr in
let sigma = Evd.merge_universe_context sigma ctx in
let t_constr_type = Retyping.get_type_of env sigma t_constr in
let tac =
@@ -701,21 +711,8 @@ END
hget_evar
*)
-let hget_evar n =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let sigma = Tacmach.New.project gl in
- let concl = Proofview.Goal.concl gl in
- let evl = evar_list concl in
- if List.length evl < n then
- error "Not enough uninstantiated existential variables.";
- if n <= 0 then error "Incorrect existential variable index.";
- let ev = List.nth evl (n-1) in
- let ev_type = existential_type sigma ev in
- change_concl (mkLetIn (Anonymous,mkEvar ev,ev_type,concl))
- end }
-
TACTIC EXTEND hget_evar
-| [ "hget_evar" int_or_var(n) ] -> [ hget_evar n ]
+| [ "hget_evar" int_or_var(n) ] -> [ Evar_tactics.hget_evar n ]
END
(**********************************************************************)
@@ -731,7 +728,7 @@ END
exception Found of unit Proofview.tactic
let rewrite_except h =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let hyps = Tacmach.New.pf_ids_of_hyps gl in
Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else
Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false))
@@ -750,11 +747,11 @@ let refl_equal =
should be replaced by a call to the tactic but I don't know how to
call it before it is defined. *)
let mkCaseEq a : unit Proofview.tactic =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let type_of_a = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g a) gl in
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let type_of_a = Tacmach.New.pf_unsafe_type_of gl a in
Tacticals.New.tclTHENLIST
- [Tactics.generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])];
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ [Tactics.generalize [(mkApp(EConstr.of_constr (delayed_force refl_equal), [| type_of_a; a|]))];
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
(** FIXME: this looks really wrong. Does anybody really use this tactic? *)
@@ -766,16 +763,16 @@ let mkCaseEq a : unit Proofview.tactic =
let case_eq_intros_rewrite x =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let n = nb_prod (Proofview.Goal.concl gl) in
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let n = nb_prod (Tacmach.New.project gl) (Proofview.Goal.concl gl) in
(* Pp.msgnl (Printer.pr_lconstr x); *)
Tacticals.New.tclTHENLIST [
mkCaseEq x;
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let hyps = Tacmach.New.pf_ids_of_hyps gl in
- let n' = nb_prod concl in
- let h = Tacmach.New.of_old (fun g -> fresh_id hyps (Id.of_string "heq") g) gl in
+ let n' = nb_prod (Tacmach.New.project gl) concl in
+ let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in
Tacticals.New.tclTHENLIST [
Tacticals.New.tclDO (n'-n-1) intro;
introduction h;
@@ -784,36 +781,37 @@ let case_eq_intros_rewrite x =
]
end }
-let rec find_a_destructable_match t =
+let rec find_a_destructable_match sigma t =
let cl = induction_arg_of_quantified_hyp (NamedHyp (Id.of_string "x")) in
let cl = [cl, (None, None), None], None in
let dest = TacAtom (Loc.ghost, TacInductionDestruct(false, false, cl)) in
- match kind_of_term t with
- | Case (_,_,x,_) when closed0 x ->
- if isVar x then
+ match EConstr.kind sigma t with
+ | Case (_,_,x,_) when closed0 sigma x ->
+ if isVar sigma x then
(* TODO check there is no rel n. *)
raise (Found (Tacinterp.eval_tactic dest))
else
(* let _ = Pp.msgnl (Printer.pr_lconstr x) in *)
raise (Found (case_eq_intros_rewrite x))
- | _ -> iter_constr find_a_destructable_match t
+ | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t
let destauto t =
- try find_a_destructable_match t;
+ Proofview.tclEVARMAP >>= fun sigma ->
+ try find_a_destructable_match sigma t;
Tacticals.New.tclZEROMSG (str "No destructable match found")
with Found tac -> tac
let destauto_in id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let ctype = Tacmach.New.of_old (fun g -> Tacmach.pf_unsafe_type_of g (mkVar id)) gl in
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in
(* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *)
(* Pp.msgnl (Printer.pr_lconstr (ctype)); *)
destauto ctype
end }
TACTIC EXTEND destauto
-| [ "destauto" ] -> [ Proofview.Goal.nf_enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ]
+| [ "destauto" ] -> [ Proofview.Goal.enter { enter = begin fun gl -> destauto (Proofview.Goal.concl gl) end } ]
| [ "destauto" "in" hyp(id) ] -> [ destauto_in id ]
END
@@ -823,8 +821,9 @@ END
let eq_constr x y =
Proofview.Goal.enter { enter = begin fun gl ->
let evd = Tacmach.New.project gl in
- if Evarutil.eq_constr_univs_test evd evd x y then Proofview.tclUNIT ()
- else Tacticals.New.tclFAIL 0 (str "Not equal")
+ match EConstr.eq_constr_universes evd x y with
+ | Some _ -> Proofview.tclUNIT ()
+ | None -> Tacticals.New.tclFAIL 0 (str "Not equal")
end }
TACTIC EXTEND constr_eq
@@ -833,21 +832,22 @@ END
TACTIC EXTEND constr_eq_nounivs
| [ "constr_eq_nounivs" constr(x) constr(y) ] -> [
- if eq_constr_nounivs x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ]
+ Proofview.tclEVARMAP >>= fun sigma ->
+ if eq_constr_nounivs sigma x y then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "Not equal") ]
END
TACTIC EXTEND is_evar
-| [ "is_evar" constr(x) ] ->
- [ Proofview.tclBIND Proofview.tclEVARMAP begin fun sigma ->
- match Evarutil.kind_of_term_upto sigma x with
- | Evar _ -> Proofview.tclUNIT ()
- | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar")
- end
+| [ "is_evar" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
+ | Evar _ -> Proofview.tclUNIT ()
+ | _ -> Tacticals.New.tclFAIL 0 (str "Not an evar")
]
END
+let has_evar sigma c =
let rec has_evar x =
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Evar _ -> true
| Rel _ | Var _ | Meta _ | Sort _ | Const _ | Ind _ | Construct _ ->
false
@@ -866,57 +866,68 @@ and has_evar_array x =
Array.exists has_evar x
and has_evar_prec (_, ts1, ts2) =
Array.exists has_evar ts1 || Array.exists has_evar ts2
+in
+has_evar c
TACTIC EXTEND has_evar
-| [ "has_evar" constr(x) ] ->
- [ if has_evar x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars") ]
+| [ "has_evar" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ if has_evar sigma x then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (str "No evars")
+]
END
TACTIC EXTEND is_hyp
-| [ "is_var" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_var" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Var _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (str "Not a variable or hypothesis") ]
END
TACTIC EXTEND is_fix
-| [ "is_fix" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_fix" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Fix _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a fix definition") ]
END;;
TACTIC EXTEND is_cofix
-| [ "is_cofix" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_cofix" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| CoFix _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a cofix definition") ]
END;;
TACTIC EXTEND is_ind
-| [ "is_ind" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_ind" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Ind _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not an (co)inductive datatype") ]
END;;
TACTIC EXTEND is_constructor
-| [ "is_constructor" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_constructor" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Construct _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constructor") ]
END;;
TACTIC EXTEND is_proj
-| [ "is_proj" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_proj" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Proj _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a primitive projection") ]
END;;
TACTIC EXTEND is_const
-| [ "is_const" constr(x) ] ->
- [ match kind_of_term x with
+| [ "is_const" constr(x) ] -> [
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma x with
| Const _ -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclFAIL 0 (Pp.str "not a constant") ]
END;;
@@ -1060,8 +1071,9 @@ END
let decompose l c =
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let to_ind c =
- if isInd c then Univ.out_punivs (destInd c)
+ if isInd sigma c then fst (destInd sigma c)
else error "not an inductive type"
in
let l = List.map to_ind l in
@@ -1075,10 +1087,14 @@ END
(** library/keys *)
VERNAC COMMAND EXTEND Declare_keys CLASSIFIED AS SIDEFF
-| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [
- let it c = snd (Constrintern.interp_open_constr (Global.env ()) Evd.empty c) in
- let k1 = Keys.constr_key (it c) in
- let k2 = Keys.constr_key (it c') in
+| [ "Declare" "Equivalent" "Keys" constr(c) constr(c') ] -> [
+ let get_key c =
+ let (evd, c) = Constrintern.interp_open_constr (Global.env ()) Evd.empty c in
+ let kind c = EConstr.kind evd c in
+ Keys.constr_key kind c
+ in
+ let k1 = get_key c in
+ let k2 = get_key c' in
match k1, k2 with
| Some k1, Some k2 -> Keys.declare_equiv_keys k1 k2
| _ -> () ]
diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4
index f75ea7087..dfa8331ff 100644
--- a/plugins/ltac/g_auto.ml4
+++ b/plugins/ltac/g_auto.ml4
@@ -16,6 +16,7 @@ open Pcoq.Constr
open Pltac
open Hints
open Tacexpr
+open Proofview.Notations
open Names
DECLARE PLUGIN "g_auto"
@@ -49,7 +50,11 @@ let eval_uconstrs ist cs =
fail_evar = false;
expand_evars = true
} in
- List.map (fun c -> Pretyping.type_uconstr ~flags ist c) cs
+ let map c = { delayed = fun env sigma ->
+ let Sigma.Sigma (c, sigma, p) = c.delayed env sigma in
+ Sigma.Sigma (c, sigma, p)
+ } in
+ List.map (fun c -> map (Pretyping.type_uconstr ~flags ist c)) cs
let pr_auto_using_raw _ _ _ = Pptactic.pr_auto_using Ppconstr.pr_constr_expr
let pr_auto_using_glob _ _ _ = Pptactic.pr_auto_using (fun (c,_) -> Printer.pr_glob_constr c)
diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4
index ca9537c82..40f30c794 100644
--- a/plugins/ltac/g_class.ml4
+++ b/plugins/ltac/g_class.ml4
@@ -89,7 +89,7 @@ TACTIC EXTEND is_ground
END
TACTIC EXTEND autoapply
- [ "autoapply" constr(c) "using" preident(i) ] -> [ Proofview.V82.tactic (autoapply c i) ]
+ [ "autoapply" constr(c) "using" preident(i) ] -> [ autoapply c i ]
END
(** TODO: DEPRECATE *)
@@ -98,18 +98,20 @@ open Term
open Proofview.Goal
open Proofview.Notations
-let rec eq_constr_mod_evars x y =
- match kind_of_term x, kind_of_term y with
+let rec eq_constr_mod_evars sigma x y =
+ let open EConstr in
+ match EConstr.kind sigma x, EConstr.kind sigma y with
| Evar (e1, l1), Evar (e2, l2) when not (Evar.equal e1 e2) -> true
- | _, _ -> compare_constr eq_constr_mod_evars x y
+ | _, _ -> compare_constr sigma (fun x y -> eq_constr_mod_evars sigma x y) x y
let progress_evars t =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let check =
- Proofview.Goal.nf_enter { enter = begin fun gl' ->
+ Proofview.Goal.enter { enter = begin fun gl' ->
+ let sigma = Tacmach.New.project gl' in
let newconcl = Proofview.Goal.concl gl' in
- if eq_constr_mod_evars concl newconcl
+ if eq_constr_mod_evars sigma concl newconcl
then Tacticals.New.tclFAIL 0 (Pp.str"No progress made (modulo evars)")
else Proofview.tclUNIT ()
end }
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index dc418d530..39ae1f41d 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -79,8 +79,8 @@ type 'a glob_extra_genarg_printer =
'a -> std_ppcmds
type 'a extra_genarg_printer =
- (Term.constr -> std_ppcmds) ->
- (Term.constr -> std_ppcmds) ->
+ (EConstr.constr -> std_ppcmds) ->
+ (EConstr.constr -> std_ppcmds) ->
(tolerability -> Val.t -> std_ppcmds) ->
'a -> std_ppcmds
@@ -1153,23 +1153,24 @@ type 'a extra_genarg_printer =
let pr_glob_tactic env = pr_glob_tactic_level env ltop
let strip_prod_binders_constr n ty =
+ let ty = EConstr.Unsafe.to_constr ty in
let rec strip_ty acc n ty =
- if n=0 then (List.rev acc, ty) else
+ 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.ghost,na],a)::acc) (n-1) b
+ strip_ty (([Loc.ghost,na],EConstr.of_constr a)::acc) (n-1) b
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
- let pr_atomic_tactic_level env n t =
+ let pr_atomic_tactic_level env sigma n t =
let prtac n (t:atomic_tactic_expr) =
let pr = {
pr_tactic = (fun _ _ -> str "<tactic>");
- pr_constr = pr_constr_env env Evd.empty;
+ pr_constr = (fun c -> pr_econstr_env env sigma c);
pr_dconstr = pr_and_constr_expr (pr_glob_constr_env env);
- pr_lconstr = pr_lconstr_env env Evd.empty;
- pr_pattern = pr_constr_pattern_env env Evd.empty;
- pr_lpattern = pr_lconstr_pattern_env env Evd.empty;
+ pr_lconstr = (fun c -> pr_leconstr_env env sigma c);
+ pr_pattern = pr_constr_pattern_env env sigma;
+ pr_lpattern = pr_lconstr_pattern_env env sigma;
pr_constant = pr_evaluable_reference_env env;
pr_reference = pr_located pr_ltac_constant;
pr_name = pr_id;
@@ -1200,7 +1201,7 @@ type 'a extra_genarg_printer =
let pr_extend pr lev ml args =
pr_extend_gen pr lev ml args
- let pr_atomic_tactic env = pr_atomic_tactic_level env ltop
+ let pr_atomic_tactic env sigma c = pr_atomic_tactic_level env sigma ltop c
let declare_extra_genarg_pprule wit
(f : 'a raw_extra_genarg_printer)
@@ -1217,7 +1218,7 @@ let declare_extra_genarg_pprule wit
in
let h x =
let env = Global.env () in
- h (pr_constr_env env Evd.empty) (pr_lconstr_env env Evd.empty) (fun _ _ -> str "<tactic>") x
+ h (pr_econstr_env env Evd.empty) (pr_leconstr_env env Evd.empty) (fun _ _ -> str "<tactic>") x
in
Genprint.register_print0 wit f g h
@@ -1247,7 +1248,7 @@ let () =
wit_intro_pattern
(Miscprint.pr_intro_pattern pr_constr_expr)
(Miscprint.pr_intro_pattern (fun (c,_) -> pr_glob_constr c))
- (Miscprint.pr_intro_pattern (fun c -> pr_constr (fst (run_delayed c))));
+ (Miscprint.pr_intro_pattern (fun c -> pr_econstr (fst (run_delayed c))));
Genprint.register_print0
wit_clause_dft_concl
(pr_clauses (Some true) pr_lident)
@@ -1258,7 +1259,7 @@ let () =
wit_constr
Ppconstr.pr_constr_expr
(fun (c, _) -> Printer.pr_glob_constr c)
- Printer.pr_constr
+ Printer.pr_econstr
;
Genprint.register_print0
wit_uconstr
@@ -1270,25 +1271,25 @@ let () =
wit_open_constr
Ppconstr.pr_constr_expr
(fun (c, _) -> Printer.pr_glob_constr c)
- Printer.pr_constr
+ Printer.pr_econstr
;
Genprint.register_print0 wit_red_expr
(pr_red_expr (pr_constr_expr, pr_lconstr_expr, pr_or_by_notation pr_reference, pr_constr_pattern_expr))
(pr_red_expr (pr_and_constr_expr pr_glob_constr, pr_and_constr_expr pr_lglob_constr, pr_or_var (pr_and_short_name pr_evaluable_reference), pr_pat_and_constr_expr pr_glob_constr))
- (pr_red_expr (pr_constr, pr_lconstr, pr_evaluable_reference, pr_constr_pattern));
+ (pr_red_expr (pr_econstr, pr_leconstr, pr_evaluable_reference, pr_constr_pattern));
Genprint.register_print0 wit_quant_hyp pr_quantified_hypothesis pr_quantified_hypothesis pr_quantified_hypothesis;
Genprint.register_print0 wit_bindings
(pr_bindings_no_with pr_constr_expr pr_lconstr_expr)
(pr_bindings_no_with (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr))
- (fun it -> pr_bindings_no_with pr_constr pr_lconstr (fst (run_delayed it)));
+ (fun it -> pr_bindings_no_with pr_econstr pr_leconstr (fst (run_delayed it)));
Genprint.register_print0 wit_constr_with_bindings
(pr_with_bindings pr_constr_expr pr_lconstr_expr)
(pr_with_bindings (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr))
- (fun it -> pr_with_bindings pr_constr pr_lconstr (fst (run_delayed it)));
+ (fun it -> pr_with_bindings pr_econstr pr_leconstr (fst (run_delayed it)));
Genprint.register_print0 Tacarg.wit_destruction_arg
(pr_destruction_arg pr_constr_expr pr_lconstr_expr)
(pr_destruction_arg (pr_and_constr_expr pr_glob_constr) (pr_and_constr_expr pr_lglob_constr))
- (fun it -> pr_destruction_arg pr_constr pr_lconstr (run_delayed_destruction_arg it));
+ (fun it -> pr_destruction_arg pr_econstr pr_leconstr (run_delayed_destruction_arg it));
Genprint.register_print0 Stdarg.wit_int int int int;
Genprint.register_print0 Stdarg.wit_bool pr_bool pr_bool pr_bool;
Genprint.register_print0 Stdarg.wit_unit pr_unit pr_unit pr_unit;
diff --git a/plugins/ltac/pptactic.mli b/plugins/ltac/pptactic.mli
index 43e22dba3..729338fb9 100644
--- a/plugins/ltac/pptactic.mli
+++ b/plugins/ltac/pptactic.mli
@@ -36,8 +36,8 @@ type 'a glob_extra_genarg_printer =
'a -> std_ppcmds
type 'a extra_genarg_printer =
- (Term.constr -> std_ppcmds) ->
- (Term.constr -> std_ppcmds) ->
+ (EConstr.t -> std_ppcmds) ->
+ (EConstr.t -> std_ppcmds) ->
(tolerability -> Val.t -> std_ppcmds) ->
'a -> std_ppcmds
@@ -100,7 +100,7 @@ val pr_raw_tactic_level : tolerability -> raw_tactic_expr -> std_ppcmds
val pr_glob_tactic : env -> glob_tactic_expr -> std_ppcmds
-val pr_atomic_tactic : env -> atomic_tactic_expr -> std_ppcmds
+val pr_atomic_tactic : env -> Evd.evar_map -> atomic_tactic_expr -> std_ppcmds
val pr_hintbases : string list option -> std_ppcmds
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index 3c5a109c0..b84be4600 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -13,6 +13,7 @@ open Util
open Nameops
open Namegen
open Term
+open EConstr
open Vars
open Reduction
open Tacticals.New
@@ -31,6 +32,7 @@ open Decl_kinds
open Elimschemes
open Environ
open Termops
+open EConstr
open Libnames
open Sigma.Notations
open Proofview.Notations
@@ -97,7 +99,7 @@ let new_cstr_evar (evd,cstrs) env t =
let evd = Sigma.Unsafe.of_evar_map evd in
let Sigma (t, evd', _) = Evarutil.new_evar ~store:s env evd t in
let evd' = Sigma.to_evar_map evd' in
- let ev, _ = destEvar t in
+ let ev, _ = destEvar evd' t in
(evd', Evar.Set.add ev cstrs), t
(** Building or looking up instances. *)
@@ -214,7 +216,7 @@ end) = struct
match obj with
| None | Some (_, None) ->
let evars, relty = mk_relation env evars ty in
- if closed0 ty then
+ if closed0 (goalevars evars) ty then
let env' = Environ.reset_with_named_context (Environ.named_context_val env) env in
new_cstr_evar evars env' relty
else new_cstr_evar evars newenv relty
@@ -222,10 +224,10 @@ end) = struct
in
let rec aux env evars ty l =
let t = Reductionops.whd_all env (goalevars evars) ty in
- match kind_of_term t, l with
+ match EConstr.kind (goalevars evars) t, l with
| Prod (na, ty, b), obj :: cstrs ->
let b = Reductionops.nf_betaiota (goalevars evars) b in
- if noccurn 1 b (* non-dependent product *) then
+ if noccurn (goalevars evars) 1 b (* non-dependent product *) then
let ty = Reductionops.nf_betaiota (goalevars evars) ty in
let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in
let evars, relty = mk_relty evars env ty obj in
@@ -233,7 +235,7 @@ end) = struct
evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs
else
let (evars, b, arg, cstrs) =
- aux (Environ.push_rel (LocalAssum (na, ty)) env) evars b cstrs
+ aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs
in
let ty = Reductionops.nf_betaiota (goalevars evars) ty in
let pred = mkLambda (na, ty, b) in
@@ -253,30 +255,30 @@ end) = struct
(** Folding/unfolding of the tactic constants. *)
- let unfold_impl t =
- match kind_of_term t with
+ let unfold_impl sigma t =
+ match EConstr.kind sigma t with
| App (arrow, [| a; b |])(* when eq_constr arrow (Lazy.force impl) *) ->
mkProd (Anonymous, a, lift 1 b)
| _ -> assert false
- let unfold_all t =
- match kind_of_term t with
+ let unfold_all sigma t =
+ match EConstr.kind sigma t with
| App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) ->
- (match kind_of_term b with
+ (match EConstr.kind sigma b with
| Lambda (n, ty, b) -> mkProd (n, ty, b)
| _ -> assert false)
| _ -> assert false
- let unfold_forall t =
- match kind_of_term t with
+ let unfold_forall sigma t =
+ match EConstr.kind sigma t with
| App (id, [| a; b |]) (* when eq_constr id (Lazy.force coq_all) *) ->
- (match kind_of_term b with
+ (match EConstr.kind sigma b with
| Lambda (n, ty, b) -> mkProd (n, ty, b)
| _ -> assert false)
| _ -> assert false
let arrow_morphism env evd ta tb a b =
- let ap = is_Prop ta and bp = is_Prop tb in
+ let ap = is_Prop (goalevars evd) ta and bp = is_Prop (goalevars evd) tb in
if ap && bp then app_poly env evd impl [| a; b |], unfold_impl
else if ap then (* Domain in Prop, CoDomain in Type *)
(app_poly env evd arrow [| a; b |]), unfold_impl
@@ -286,28 +288,28 @@ end) = struct
else (* None in Prop, use arrow *)
(app_poly env evd arrow [| a; b |]), unfold_impl
- let rec decomp_pointwise n c =
+ let rec decomp_pointwise sigma n c =
if Int.equal n 0 then c
else
- match kind_of_term c with
- | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f ->
- decomp_pointwise (pred n) relb
- | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f ->
- decomp_pointwise (pred n) (Reductionops.beta_applist (arelb, [mkRel 1]))
+ match EConstr.kind sigma c with
+ | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f ->
+ decomp_pointwise sigma (pred n) relb
+ | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f ->
+ decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (arelb, [mkRel 1]))
| _ -> invalid_arg "decomp_pointwise"
- let rec apply_pointwise rel = function
+ let rec apply_pointwise sigma rel = function
| arg :: args ->
- (match kind_of_term rel with
- | App (f, [| a; b; relb |]) when Globnames.is_global (pointwise_relation_ref ()) f ->
- apply_pointwise relb args
- | App (f, [| a; b; arelb |]) when Globnames.is_global (forall_relation_ref ()) f ->
- apply_pointwise (Reductionops.beta_applist (arelb, [arg])) args
+ (match EConstr.kind sigma rel with
+ | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f ->
+ apply_pointwise sigma relb args
+ | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f ->
+ apply_pointwise sigma (Reductionops.beta_applist sigma (arelb, [arg])) args
| _ -> invalid_arg "apply_pointwise")
| [] -> rel
let pointwise_or_dep_relation env evd n t car rel =
- if noccurn 1 car && noccurn 1 rel then
+ if noccurn (goalevars evd) 1 car && noccurn (goalevars evd) 1 rel then
app_poly env evd pointwise_relation [| t; lift (-1) car; lift (-1) rel |]
else
app_poly env evd forall_relation
@@ -324,14 +326,15 @@ end) = struct
let rec aux evars env prod n =
if Int.equal n 0 then start evars env prod
else
- match kind_of_term (Reduction.whd_all env prod) with
+ let sigma = goalevars evars in
+ match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with
| Prod (na, ty, b) ->
- if noccurn 1 b then
+ if noccurn sigma 1 b then
let b' = lift (-1) b in
let evars, rb = aux evars env b' (pred n) in
app_poly env evars pointwise_relation [| ty; b'; rb |]
else
- let evars, rb = aux evars (Environ.push_rel (LocalAssum (na, ty)) env) b (pred n) in
+ let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in
app_poly env evars forall_relation
[| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |]
| _ -> raise Not_found
@@ -342,24 +345,25 @@ end) = struct
try let evars, found = aux evars env ty (succ (List.length args)) in
Some (evars, found, c, ty, arg :: args)
with Not_found ->
- let ty = whd_all env ty in
- find env (mkApp (c, [| arg |])) (prod_applist ty [arg]) args
+ let sigma = goalevars evars in
+ let ty = Reductionops.whd_all env sigma ty in
+ find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args
in find env c ty args
let unlift_cstr env sigma = function
| None -> None
- | Some codom -> Some (decomp_pointwise 1 codom)
+ | Some codom -> Some (decomp_pointwise (goalevars sigma) 1 codom)
(** Looking up declared rewrite relations (instances of [RewriteRelation]) *)
let is_applied_rewrite_relation env sigma rels t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| App (c, args) when Array.length args >= 2 ->
- let head = if isApp c then fst (destApp c) else c in
- if Globnames.is_global (coq_eq_ref ()) head then None
+ let head = if isApp sigma c then fst (destApp sigma c) else c in
+ if Termops.is_global sigma (coq_eq_ref ()) head then None
else
(try
let params, args = Array.chop (Array.length args - 2) args in
- let env' = Environ.push_rel_context rels env in
+ let env' = push_rel_context rels env in
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma ((evar, _), evars, _) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in
let evars = Sigma.to_evar_map evars in
@@ -430,7 +434,7 @@ module TypeGlobal = struct
end
let sort_of_rel env evm rel =
- Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel)
+ ESorts.kind evm (Reductionops.sort_of_arity env evm (Retyping.get_type_of env evm rel))
let is_applied_rewrite_relation = PropGlobal.is_applied_rewrite_relation
@@ -480,7 +484,7 @@ let error_no_relation () = error "Cannot find a relation to rewrite."
let rec decompose_app_rel env evd t =
(** Head normalize for compatibility with the old meta mechanism *)
let t = Reductionops.whd_betaiota evd t in
- match kind_of_term t with
+ match EConstr.kind evd t with
| App (f, [||]) -> assert false
| App (f, [|arg|]) ->
let (f', argl, argr) = decompose_app_rel env evd arg in
@@ -499,7 +503,7 @@ let rec decompose_app_rel env evd t =
let decompose_app_rel env evd t =
let (rel, t1, t2) = decompose_app_rel env evd t in
let ty = Retyping.get_type_of env evd rel in
- let () = if not (Reduction.is_arity env ty) then error_no_relation () in
+ let () = if not (Reductionops.is_arity env evd ty) then error_no_relation () in
(rel, t1, t2)
let decompose_applied_relation env sigma (c,l) =
@@ -617,9 +621,10 @@ let solve_remaining_by env sigma holes by =
| Some tac ->
let map h =
if h.Clenv.hole_deps then None
- else
- let (evk, _) = destEvar (h.Clenv.hole_evar) in
+ else match EConstr.kind sigma h.Clenv.hole_evar with
+ | Evar (evk, _) ->
Some evk
+ | _ -> None
in
(** Only solve independent holes *)
let indep = List.map_filter map holes in
@@ -639,7 +644,7 @@ let solve_remaining_by env sigma holes by =
(** Evar should not be defined, but just in case *)
| Some evi ->
let env = Environ.reset_with_named_context evi.evar_hyps env in
- let ty = evi.evar_concl in
+ let ty = EConstr.of_constr evi.evar_concl in
let c, sigma = Pfedit.refine_by_tactic env sigma ty solve_tac in
Evd.define evk c sigma
in
@@ -714,7 +719,7 @@ let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs)
let sigma = Typeclasses.resolve_typeclasses ~filter:(no_constraints cstrs)
~fail:true env sigma in
let evd = solve_remaining_by env sigma holes by in
- let nf c = Evarutil.nf_evar evd (Reductionops.nf_meta evd c) in
+ let nf c = Reductionops.nf_evar evd (Reductionops.nf_meta evd c) in
let c1 = nf c1 and c2 = nf c2
and rew_car = nf car and rel = nf rel
and prf = nf prf in
@@ -754,9 +759,9 @@ let default_flags = { under_lambdas = true; on_morphisms = true; }
let get_opt_rew_rel = function RewPrf (rel, prf) -> Some rel | _ -> None
let make_eq () =
-(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ()))
let make_eq_refl () =
-(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq_refl ())
+(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq_refl ()))
let get_rew_prf r = match r.rew_prf with
| RewPrf (rel, prf) -> rel, prf
@@ -769,7 +774,7 @@ let poly_subrelation sort =
if sort then PropGlobal.subrelation else TypeGlobal.subrelation
let resolve_subrelation env avoid car rel sort prf rel' res =
- if eq_constr rel rel' then res
+ if Termops.eq_constr (fst res.rew_evars) rel rel' then res
else
let evars, app = app_poly_check env res.rew_evars (poly_subrelation sort) [|car; rel; rel'|] in
let evars, subrel = new_cstr_evar evars env app in
@@ -805,7 +810,7 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev
if b then PropGlobal.do_subrelation, PropGlobal.apply_subrelation
else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation
in
- Environ.push_named
+ EConstr.push_named
(LocalDef (Id.of_string "do_subrelation",
snd (app_poly_sort b env evars dosub [||]),
snd (app_poly_nocheck env evars appsub [||])))
@@ -837,8 +842,8 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev
x :: acc, x :: subst, evars, sigargs, x :: typeargs')
([], [], evars, sigargs, []) args args'
in
- let proof = applistc proj (List.rev projargs) in
- let newt = applistc m' (List.rev typeargs) in
+ let proof = applist (proj, List.rev projargs) in
+ let newt = applist (m', List.rev typeargs) in
match respars with
[ a, Some r ] -> evars, proof, substl subst a, substl subst r, oldt, fnewt newt
| _ -> assert(false)
@@ -861,13 +866,13 @@ let apply_rule unify loccs : int pure_strategy =
in
{ strategy = fun { state = occ ; env ; unfresh ;
term1 = t ; ty1 = ty ; cstr ; evars } ->
- let unif = if isEvar t then None else unify env evars t in
+ let unif = if isEvar (goalevars evars) t then None else unify env evars t in
match unif with
| None -> (occ, Fail)
| Some rew ->
let occ = succ occ in
if not (is_occ occ) then (occ, Fail)
- else if eq_constr t rew.rew_to then (occ, Identity)
+ else if Termops.eq_constr (fst rew.rew_evars) t rew.rew_to then (occ, Identity)
else
let res = { rew with rew_car = ty } in
let rel, prf = get_rew_prf res in
@@ -921,17 +926,17 @@ let reset_env env =
Environ.push_rel_context (Environ.rel_context env) env'
let fold_match ?(force=false) env sigma c =
- let (ci, p, c, brs) = destCase c in
+ let (ci, p, c, brs) = destCase sigma c in
let cty = Retyping.get_type_of env sigma c in
let dep, pred, exists, (sk,eff) =
let env', ctx, body =
- let ctx, pred = decompose_lam_assum p in
- let env' = Environ.push_rel_context ctx env in
+ let ctx, pred = decompose_lam_assum sigma p in
+ let env' = push_rel_context ctx env in
env', ctx, pred
in
let sortp = Retyping.get_sort_family_of env' sigma body in
let sortc = Retyping.get_sort_family_of env sigma cty in
- let dep = not (noccurn 1 body) in
+ let dep = not (noccurn sigma 1 body) in
let pred = if dep then p else
it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx)
in
@@ -955,7 +960,7 @@ let fold_match ?(force=false) env sigma c =
else raise Not_found
in
let app =
- let ind, args = Inductive.find_rectype env cty in
+ let ind, args = Inductiveops.find_mrectype env sigma cty in
let pars, args = List.chop ci.ci_npar args in
let meths = List.map (fun br -> br) (Array.to_list brs) in
applist (mkConst sk, pars @ [pred] @ meths @ args @ [c])
@@ -963,9 +968,10 @@ let fold_match ?(force=false) env sigma c =
sk, (if exists then env else reset_env env), app, eff
let unfold_match env sigma sk app =
- match kind_of_term app with
- | App (f', args) when eq_constant (fst (destConst f')) sk ->
+ match EConstr.kind sigma app with
+ | App (f', args) when eq_constant (fst (destConst sigma f')) sk ->
let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in
+ let v = EConstr.of_constr v in
Reductionops.whd_beta sigma (mkApp (v, args))
| _ -> app
@@ -975,7 +981,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
let rec aux { state ; env ; unfresh ;
term1 = t ; ty1 = ty ; cstr = (prop, cstr) ; evars } =
let cstr' = Option.map (fun c -> (ty, Some c)) cstr in
- match kind_of_term t with
+ match EConstr.kind (goalevars evars) t with
| App (m, args) ->
let rewrite_args state success =
let state, (args', evars', progress) =
@@ -1055,7 +1061,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
let app = if prop then PropGlobal.apply_pointwise
else TypeGlobal.apply_pointwise
in
- RewPrf (app rel argsl, mkApp (prf, args))
+ RewPrf (app (goalevars evars) rel argsl, mkApp (prf, args))
| x -> x
in
let res =
@@ -1072,9 +1078,9 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
in state, res
else rewrite_args state None
- | Prod (n, x, b) when noccurn 1 b ->
+ | Prod (n, x, b) when noccurn (goalevars evars) 1 b ->
let b = subst1 mkProp b in
- let tx = Retyping.get_type_of env (goalevars evars) x
+ let tx = Retyping.get_type_of env (goalevars evars) x
and tb = Retyping.get_type_of env (goalevars evars) b in
let arr = if prop then PropGlobal.arrow_morphism
else TypeGlobal.arrow_morphism
@@ -1085,7 +1091,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
cstr = (prop,cstr) ; evars = evars' } in
let res =
match res with
- | Success r -> Success { r with rew_to = unfold r.rew_to }
+ | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to }
| Fail | Identity -> res
in state, res
@@ -1106,7 +1112,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
| Prod (n, dom, codom) ->
let lam = mkLambda (n, dom, codom) in
let (evars', app), unfold =
- if eq_constr ty mkProp then
+ if eq_constr (fst evars) ty mkProp then
(app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all
else
let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in
@@ -1117,7 +1123,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
cstr = (prop,cstr) ; evars = evars' } in
let res =
match res with
- | Success r -> Success { r with rew_to = unfold r.rew_to }
+ | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to }
| Fail | Identity -> res
in state, res
@@ -1152,7 +1158,7 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy =
| Lambda (n, t, b) when flags.under_lambdas ->
let n' = name_app (fun id -> Tactics.fresh_id_in_env unfresh id env) n in
let open Context.Rel.Declaration in
- let env' = Environ.push_rel (LocalAssum (n', t)) env in
+ let env' = EConstr.push_rel (LocalAssum (n', t)) env in
let bty = Retyping.get_type_of env' (goalevars evars) b in
let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in
let state, b' = s.strategy { state ; env = env' ; unfresh ;
@@ -1381,7 +1387,7 @@ module Strategies =
let inj_open hint = (); fun sigma ->
let ctx = Evd.evar_universe_context_of hint.Autorewrite.rew_ctx in
let sigma = Evd.merge_universe_context sigma ctx in
- (sigma, (hint.Autorewrite.rew_lemma, NoBindings))
+ (sigma, (EConstr.of_constr hint.Autorewrite.rew_lemma, NoBindings))
let old_hints (db : string) : 'a pure_strategy =
let rules = Autorewrite.find_rewrites db in
@@ -1391,6 +1397,7 @@ module Strategies =
let hints (db : string) : 'a pure_strategy = { strategy =
fun ({ term1 = t } as input) ->
+ let t = EConstr.Unsafe.to_constr t in
let rules = Autorewrite.find_matches db t in
let lemma hint = (inj_open hint, hint.Autorewrite.rew_l2r,
hint.Autorewrite.rew_tac) in
@@ -1404,7 +1411,7 @@ module Strategies =
let sigma = Sigma.Unsafe.of_evar_map (goalevars evars) in
let Sigma (t', sigma, _) = rfn.Reductionops.e_redfun env sigma t in
let evars' = Sigma.to_evar_map sigma in
- if eq_constr t' t then
+ if Termops.eq_constr evars' t' t then
state, Identity
else
state, Success { rew_car = ty; rew_from = t; rew_to = t';
@@ -1423,7 +1430,7 @@ module Strategies =
in
try
let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in
- let c' = Evarutil.nf_evar sigma c in
+ let c' = Reductionops.nf_evar sigma c in
state, Success { rew_car = ty; rew_from = t; rew_to = c';
rew_prf = RewCast DEFAULTcast;
rew_evars = (sigma, snd evars) }
@@ -1496,7 +1503,7 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul
| Success res ->
let (_, cstrs) = res.rew_evars in
let evars' = solve_constraints env res.rew_evars in
- let newt = Evarutil.nf_evar evars' res.rew_to in
+ let newt = Reductionops.nf_evar evars' res.rew_to in
let evars = (* Keep only original evars (potentially instantiated) and goal evars,
the rest has been defined and substituted already. *)
Evar.Set.fold
@@ -1504,20 +1511,20 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul
if not (Evd.is_defined acc ev) then
user_err ~hdr:"rewrite"
(str "Unsolved constraint remaining: " ++ spc () ++
- Evd.pr_evar_info (Evd.find acc ev))
+ Termops.pr_evar_info (Evd.find acc ev))
else Evd.remove acc ev)
cstrs evars'
in
let res = match res.rew_prf with
| RewCast c -> None
| RewPrf (rel, p) ->
- let p = nf_zeta env evars' (Evarutil.nf_evar evars' p) in
+ let p = nf_zeta env evars' (Reductionops.nf_evar evars' p) in
let term =
match abs with
| None -> p
| Some (t, ty) ->
- let t = Evarutil.nf_evar evars' t in
- let ty = Evarutil.nf_evar evars' ty in
+ let t = Reductionops.nf_evar evars' t in
+ let ty = Reductionops.nf_evar evars' ty in
mkApp (mkLambda (Name (Id.of_string "lemma"), ty, p), [| t |])
in
let proof = match is_hyp with
@@ -1527,23 +1534,24 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul
in Some (Some (evars, res, newt))
(** Insert a declaration after the last declaration it depends on *)
-let rec insert_dependent env decl accu hyps = match hyps with
+let rec insert_dependent env sigma decl accu hyps = match hyps with
| [] -> List.rev_append accu [decl]
| ndecl :: rem ->
- if occur_var_in_decl env (NamedDecl.get_id ndecl) decl then
+ if occur_var_in_decl env sigma (NamedDecl.get_id ndecl) decl then
List.rev_append accu (decl :: hyps)
else
- insert_dependent env decl (ndecl :: accu) rem
+ insert_dependent env sigma decl (ndecl :: accu) rem
let assert_replacing id newt tac =
- let prf = Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let prf = Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
- let ctx = Environ.named_context env in
+ let sigma = Tacmach.New.project gl in
+ let ctx = named_context env in
let after, before = List.split_when (NamedDecl.get_id %> Id.equal id) ctx in
let nc = match before with
| [] -> assert false
- | d :: rem -> insert_dependent env (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem
+ | d :: rem -> insert_dependent env sigma (LocalAssum (NamedDecl.get_id d, newt)) [] after @ rem
in
let env' = Environ.reset_with_named_context (val_of_named_context nc) env in
Refine.refine ~unsafe:false { run = begin fun sigma ->
@@ -1553,7 +1561,7 @@ let assert_replacing id newt tac =
let n = NamedDecl.get_id d in
if Id.equal n id then ev' else mkVar n
in
- let (e, _) = destEvar ev in
+ let (e, _) = destEvar (Sigma.to_evar_map sigma) ev in
Sigma (mkEvar (e, Array.map_of_list map nc), sigma, p +> q)
end }
end } in
@@ -1603,22 +1611,22 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause =
Proofview.Unsafe.tclEVARS undef <*>
convert_concl_no_check newt DEFAULTcast
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let ty = match clause with
| None -> concl
- | Some id -> Environ.named_type id env
+ | Some id -> EConstr.of_constr (Environ.named_type id env)
in
let env = match clause with
| None -> env
| Some id ->
(** Only consider variables not depending on [id] *)
- let ctx = Environ.named_context env in
- let filter decl = not (occur_var_in_decl env id decl) in
+ let ctx = named_context env in
+ let filter decl = not (occur_var_in_decl env sigma id decl) in
let nctx = List.filter filter ctx in
- Environ.reset_with_named_context (Environ.val_of_named_context nctx) env
+ Environ.reset_with_named_context (val_of_named_context nctx) env
in
try
let res =
@@ -1853,9 +1861,10 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans =
let cHole = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None)
-let proper_projection r ty =
- let ctx, inst = decompose_prod_assum ty in
- let mor, args = destApp inst in
+let proper_projection sigma r ty =
+ let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i)) in
+ let ctx, inst = decompose_prod_assum sigma ty in
+ let mor, args = destApp sigma inst in
let instarg = mkApp (r, rel_vect 0 (List.length ctx)) in
let app = mkApp (Lazy.force PropGlobal.proper_proj,
Array.append args [| instarg |]) in
@@ -1866,31 +1875,34 @@ let declare_projection n instance_id r =
let env = Global.env () in
let sigma = Evd.from_env env in
let sigma,c = Evd.fresh_global env sigma r in
+ let c = EConstr.of_constr c in
let ty = Retyping.get_type_of env sigma c in
- let term = proper_projection c ty in
+ let term = proper_projection sigma c ty in
let sigma, typ = Typing.type_of env sigma term in
- let ctx, typ = decompose_prod_assum typ in
+ let ctx, typ = decompose_prod_assum sigma typ in
let typ =
let n =
let rec aux t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| App (f, [| a ; a' ; rel; rel' |])
- when Globnames.is_global (PropGlobal.respectful_ref ()) f ->
+ when Termops.is_global sigma (PropGlobal.respectful_ref ()) f ->
succ (aux rel')
| _ -> 0
in
let init =
- match kind_of_term typ with
- App (f, args) when Globnames.is_global (PropGlobal.respectful_ref ()) f ->
+ match EConstr.kind sigma typ with
+ App (f, args) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f ->
mkApp (f, fst (Array.chop (Array.length args - 2) args))
| _ -> typ
in aux init
in
- let ctx,ccl = Reductionops.splay_prod_n (Global.env()) Evd.empty (3 * n) typ
+ let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ
in it_mkProd_or_LetIn ccl ctx
in
let typ = it_mkProd_or_LetIn typ ctx in
let pl, ctx = Evd.universe_context sigma in
+ let typ = EConstr.to_constr sigma typ in
+ let term = EConstr.to_constr sigma term in
let cst =
Declare.definition_entry ~types:typ ~poly ~univs:ctx term
in
@@ -1899,11 +1911,12 @@ let declare_projection n instance_id r =
let build_morphism_signature env sigma m =
let m,ctx = Constrintern.interp_constr env sigma m in
+ let m = EConstr.of_constr m in
let sigma = Evd.from_ctx ctx in
let t = Typing.unsafe_type_of env sigma m in
let cstrs =
let rec aux t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (na, a, b) ->
None :: aux b
| _ -> []
@@ -1923,8 +1936,8 @@ let build_morphism_signature env sigma m =
let morph = e_app_poly env evd PropGlobal.proper_type [| t; sig_; m |] in
let evd = solve_constraints env !evd in
let evd = Evd.nf_constraints evd in
- let m = Evarutil.nf_evars_universes evd morph in
- Pretyping.check_evars env Evd.empty evd m;
+ let m = Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr morph) in
+ Pretyping.check_evars env Evd.empty evd (EConstr.of_constr m);
Evd.evar_universe_context evd, m
let default_morphism sign m =
@@ -1936,7 +1949,7 @@ let default_morphism sign m =
in
let evars, morph = app_poly_check env evars PropGlobal.proper_type [| t; sign; m |] in
let evars, mor = resolve_one_typeclass env (goalevars evars) morph in
- mor, proper_projection mor morph
+ mor, proper_projection sigma mor morph
let add_setoid global binders a aeq t n =
init_setoid ();
@@ -1991,7 +2004,7 @@ let add_morphism_infer glob m n =
let hook = Lemmas.mk_hook hook in
Flags.silently
(fun () ->
- Lemmas.start_proof instance_id kind (Evd.from_ctx uctx) instance hook;
+ Lemmas.start_proof instance_id kind (Evd.from_ctx uctx) (EConstr.of_constr instance) hook;
ignore (Pfedit.by (Tacinterp.interp tac))) ()
let add_morphism glob binders m s n =
@@ -2052,7 +2065,7 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env =
~flags:rewrite_conv_unif_flags
env sigma ((if l2r then c1 else c2),but)
in
- let nf c = Evarutil.nf_evar sigma c in
+ let nf c = Reductionops.nf_evar sigma c in
let c1 = if l2r then nf c' else nf c1
and c2 = if l2r then nf c2 else nf c'
and car = nf car and rel = nf rel in
@@ -2071,7 +2084,7 @@ let get_hyp gl (c,l) clause l2r =
let sigma, hi = decompose_applied_relation env evars (c,l) in
let but = match clause with
| Some id -> Tacmach.New.pf_get_hyp_typ id gl
- | None -> Evarutil.nf_evar evars (Tacmach.New.pf_concl gl)
+ | None -> Reductionops.nf_evar evars (Tacmach.New.pf_concl gl)
in
unification_rewrite l2r hi.c1 hi.c2 sigma hi.prf hi.car hi.rel but env
@@ -2082,7 +2095,7 @@ let general_rewrite_flags = { under_lambdas = false; on_morphisms = true }
(** Setoid rewriting when called with "rewrite" *)
let general_s_rewrite cl l2r occs (c,l) ~new_goals =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let abs, evd, res, sort = get_hyp gl (c,l) cl l2r in
let unify env evars t = unify_abs res l2r sort env evars t in
let app = apply_rule unify occs in
@@ -2110,13 +2123,13 @@ let _ = Hook.set Equality.general_setoid_rewrite_clause general_s_rewrite
(** [setoid_]{reflexivity,symmetry,transitivity} tactics *)
-let not_declared env ty rel =
+let not_declared env sigma ty rel =
tclFAIL 0
- (str" The relation " ++ Printer.pr_constr_env env Evd.empty rel ++ str" is not a declared " ++
+ (str" The relation " ++ Printer.pr_econstr_env env sigma rel ++ str" is not a declared " ++
str ty ++ str" relation. Maybe you need to require the Coq.Classes.RelationClasses library")
let setoid_proof ty fn fallback =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
@@ -2125,7 +2138,7 @@ let setoid_proof ty fn fallback =
try
let rel, _, _ = decompose_app_rel env sigma concl in
let (sigma, t) = Typing.type_of env sigma rel in
- let car = RelDecl.get_type (List.hd (fst (Reduction.dest_prod env t))) in
+ let car = snd (List.hd (fst (Reductionops.splay_prod env sigma t))) in
(try init_relation_classes () with _ -> raise Not_found);
fn env sigma car rel
with e -> Proofview.tclZERO e
@@ -2139,7 +2152,7 @@ let setoid_proof ty fn fallback =
begin match e with
| (Not_found, _) ->
let rel, _, _ = decompose_app_rel env sigma concl in
- not_declared env ty rel
+ not_declared env sigma ty rel
| (e, info) -> Proofview.tclZERO ~info e
end
| e' -> Proofview.tclZERO ~info e'
@@ -2185,9 +2198,10 @@ let setoid_transitivity c =
let setoid_symmetry_in id =
Proofview.V82.tactic (fun gl ->
+ let sigma = project gl in
let ctype = pf_unsafe_type_of gl (mkVar id) in
- let binders,concl = decompose_prod_assum ctype in
- let (equiv, args) = decompose_app concl in
+ let binders,concl = decompose_prod_assum sigma ctype in
+ let (equiv, args) = decompose_app sigma concl in
let rec split_last_two = function
| [c1;c2] -> [],(c1, c2)
| x::y::z -> let l,res = split_last_two (y::z) in x::l, res
diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli
index 4fdce0c84..7a20838a2 100644
--- a/plugins/ltac/rewrite.mli
+++ b/plugins/ltac/rewrite.mli
@@ -9,6 +9,7 @@
open Names
open Constr
open Environ
+open EConstr
open Constrexpr
open Tacexpr
open Misctypes
@@ -74,7 +75,7 @@ val cl_rewrite_clause :
bool -> Locus.occurrences -> Id.t option -> unit Proofview.tactic
val is_applied_rewrite_relation :
- env -> evar_map -> Context.Rel.t -> constr -> types option
+ env -> evar_map -> rel_context -> constr -> types option
val declare_relation :
?binders:local_binder_expr list -> constr_expr -> constr_expr -> Id.t ->
@@ -112,6 +113,6 @@ val apply_strategy :
strategy ->
Environ.env ->
Names.Id.t list ->
- Term.constr ->
- bool * Term.constr ->
+ constr ->
+ bool * constr ->
evars -> rewrite_result
diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml
index df38a42cb..b76009c99 100644
--- a/plugins/ltac/taccoerce.ml
+++ b/plugins/ltac/taccoerce.ml
@@ -9,6 +9,7 @@
open Util
open Names
open Term
+open EConstr
open Pattern
open Misctypes
open Genarg
@@ -17,7 +18,7 @@ open Geninterp
exception CannotCoerceTo of string
-let (wit_constr_context : (Empty.t, Empty.t, constr) Genarg.genarg_type) =
+let (wit_constr_context : (Empty.t, Empty.t, EConstr.constr) Genarg.genarg_type) =
let wit = Genarg.create_arg "constr_context" in
let () = register_val0 wit None in
wit
@@ -96,7 +97,7 @@ let is_variable env id =
(* Transforms an id into a constr if possible, or fails with Not_found *)
let constr_of_id env id =
- Term.mkVar (let _ = Environ.lookup_named id env in id)
+ EConstr.mkVar (let _ = Environ.lookup_named id env in id)
(* Gives the constr corresponding to a Constr_context tactic_arg *)
let coerce_to_constr_context v =
@@ -106,7 +107,7 @@ let coerce_to_constr_context v =
else raise (CannotCoerceTo "a term context")
(* Interprets an identifier which must be fresh *)
-let coerce_var_to_ident fresh env v =
+let coerce_var_to_ident fresh env sigma v =
let v = Value.normalize v in
let fail () = raise (CannotCoerceTo "a fresh identifier") in
if has_type v (topwit wit_intro_pattern) then
@@ -119,15 +120,16 @@ let coerce_var_to_ident fresh env v =
| None -> fail ()
| Some c ->
(* We need it fresh for intro e.g. in "Tac H = clear H; intro H" *)
- if isVar c && not (fresh && is_variable env (destVar c)) then
- destVar c
+ if isVar sigma c && not (fresh && is_variable env (destVar sigma c)) then
+ destVar sigma c
else fail ()
(* Interprets, if possible, a constr to an identifier which may not
be fresh but suitable to be given to the fresh tactic. Works for
vars, constants, inductive, constructors and sorts. *)
-let coerce_to_ident_not_fresh g env v =
+let coerce_to_ident_not_fresh env sigma v =
+let g = sigma in
let id_of_name = function
| Names.Anonymous -> Id.of_string "x"
| Names.Name x -> x in
@@ -143,7 +145,7 @@ let id_of_name = function
match Value.to_constr v with
| None -> fail ()
| Some c ->
- match Constr.kind c with
+ match EConstr.kind sigma c with
| Var id -> id
| Meta m -> id_of_name (Evd.meta_name g m)
| Evar (kn,_) ->
@@ -162,14 +164,14 @@ let id_of_name = function
basename
| Sort s ->
begin
- match s with
+ match ESorts.kind sigma s with
| Prop _ -> Label.to_id (Label.make "Prop")
| Type _ -> Label.to_id (Label.make "Type")
end
| _ -> fail()
-let coerce_to_intro_pattern env v =
+let coerce_to_intro_pattern env sigma v =
let v = Value.normalize v in
if has_type v (topwit wit_intro_pattern) then
snd (out_gen (topwit wit_intro_pattern) v)
@@ -177,14 +179,14 @@ let coerce_to_intro_pattern env v =
let id = out_gen (topwit wit_var) v in
IntroNaming (IntroIdentifier id)
else match Value.to_constr v with
- | Some c when isVar c ->
+ | Some c when isVar sigma c ->
(* This happens e.g. in definitions like "Tac H = clear H; intro H" *)
(* but also in "destruct H as (H,H')" *)
- IntroNaming (IntroIdentifier (destVar c))
+ IntroNaming (IntroIdentifier (destVar sigma c))
| _ -> raise (CannotCoerceTo "an introduction pattern")
-let coerce_to_intro_pattern_naming env v =
- match coerce_to_intro_pattern env v with
+let coerce_to_intro_pattern_naming env sigma v =
+ match coerce_to_intro_pattern env sigma v with
| IntroNaming pat -> pat
| _ -> raise (CannotCoerceTo "a naming introduction pattern")
@@ -232,7 +234,7 @@ let coerce_to_closed_constr env v =
let () = if not (List.is_empty ids) then raise (CannotCoerceTo "a term") in
c
-let coerce_to_evaluable_ref env v =
+let coerce_to_evaluable_ref env sigma v =
let fail () = raise (CannotCoerceTo "an evaluable reference") in
let v = Value.normalize v in
let ev =
@@ -253,8 +255,8 @@ let coerce_to_evaluable_ref env v =
| IndRef _ | ConstructRef _ -> fail ()
else
match Value.to_constr v with
- | Some c when isConst c -> EvalConstRef (Univ.out_punivs (destConst c))
- | Some c when isVar c -> EvalVarRef (destVar c)
+ | Some c when isConst sigma c -> EvalConstRef (fst (destConst sigma c))
+ | Some c when isVar sigma c -> EvalVarRef (destVar sigma c)
| _ -> fail ()
in if Tacred.is_evaluable env ev then ev else fail ()
@@ -266,14 +268,14 @@ let coerce_to_constr_list env v =
List.map map l
| None -> raise (CannotCoerceTo "a term list")
-let coerce_to_intro_pattern_list loc env v =
+let coerce_to_intro_pattern_list loc env sigma v =
match Value.to_list v with
| None -> raise (CannotCoerceTo "an intro pattern list")
| Some l ->
- let map v = (loc, coerce_to_intro_pattern env v) in
+ let map v = (loc, coerce_to_intro_pattern env sigma v) in
List.map map l
-let coerce_to_hyp env v =
+let coerce_to_hyp env sigma v =
let fail () = raise (CannotCoerceTo "a variable") in
let v = Value.normalize v in
if has_type v (topwit wit_intro_pattern) then
@@ -284,31 +286,31 @@ let coerce_to_hyp env v =
let id = out_gen (topwit wit_var) v in
if is_variable env id then id else fail ()
else match Value.to_constr v with
- | Some c when isVar c -> destVar c
+ | Some c when isVar sigma c -> destVar sigma c
| _ -> fail ()
-let coerce_to_hyp_list env v =
+let coerce_to_hyp_list env sigma v =
let v = Value.to_list v in
match v with
| Some l ->
- let map n = coerce_to_hyp env n in
+ let map n = coerce_to_hyp env sigma n in
List.map map l
| None -> raise (CannotCoerceTo "a variable list")
(* Interprets a qualified name *)
-let coerce_to_reference env v =
+let coerce_to_reference env sigma v =
let v = Value.normalize v in
match Value.to_constr v with
| Some c ->
begin
- try Globnames.global_of_constr c
+ try fst (Termops.global_of_constr sigma c)
with Not_found -> raise (CannotCoerceTo "a reference")
end
| None -> raise (CannotCoerceTo "a reference")
(* Quantified named or numbered hypothesis or hypothesis in context *)
(* (as in Inversion) *)
-let coerce_to_quantified_hypothesis v =
+let coerce_to_quantified_hypothesis sigma v =
let v = Value.normalize v in
if has_type v (topwit wit_intro_pattern) then
let v = out_gen (topwit wit_intro_pattern) v in
@@ -321,17 +323,17 @@ let coerce_to_quantified_hypothesis v =
else if has_type v (topwit wit_int) then
AnonHyp (out_gen (topwit wit_int) v)
else match Value.to_constr v with
- | Some c when isVar c -> NamedHyp (destVar c)
+ | Some c when isVar sigma c -> NamedHyp (destVar sigma c)
| _ -> raise (CannotCoerceTo "a quantified hypothesis")
(* Quantified named or numbered hypothesis or hypothesis in context *)
(* (as in Inversion) *)
-let coerce_to_decl_or_quant_hyp env v =
+let coerce_to_decl_or_quant_hyp env sigma v =
let v = Value.normalize v in
if has_type v (topwit wit_int) then
AnonHyp (out_gen (topwit wit_int) v)
else
- try coerce_to_quantified_hypothesis v
+ try coerce_to_quantified_hypothesis sigma v
with CannotCoerceTo _ ->
raise (CannotCoerceTo "a declared or quantified hypothesis")
diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli
index 0b67f8726..9c4ac5265 100644
--- a/plugins/ltac/taccoerce.mli
+++ b/plugins/ltac/taccoerce.mli
@@ -9,6 +9,7 @@
open Util
open Names
open Term
+open EConstr
open Misctypes
open Pattern
open Genarg
@@ -50,14 +51,14 @@ end
val coerce_to_constr_context : Value.t -> constr
-val coerce_var_to_ident : bool -> Environ.env -> Value.t -> Id.t
+val coerce_var_to_ident : bool -> Environ.env -> Evd.evar_map -> Value.t -> Id.t
-val coerce_to_ident_not_fresh : Evd.evar_map -> Environ.env -> Value.t -> Id.t
+val coerce_to_ident_not_fresh : Environ.env -> Evd.evar_map -> Value.t -> Id.t
-val coerce_to_intro_pattern : Environ.env -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr
+val coerce_to_intro_pattern : Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.delayed_open_constr intro_pattern_expr
val coerce_to_intro_pattern_naming :
- Environ.env -> Value.t -> intro_pattern_naming_expr
+ Environ.env -> Evd.evar_map -> Value.t -> intro_pattern_naming_expr
val coerce_to_hint_base : Value.t -> string
@@ -70,27 +71,27 @@ val coerce_to_uconstr : Environ.env -> Value.t -> Glob_term.closed_glob_constr
val coerce_to_closed_constr : Environ.env -> Value.t -> constr
val coerce_to_evaluable_ref :
- Environ.env -> Value.t -> evaluable_global_reference
+ Environ.env -> Evd.evar_map -> Value.t -> evaluable_global_reference
val coerce_to_constr_list : Environ.env -> Value.t -> constr list
val coerce_to_intro_pattern_list :
- Loc.t -> Environ.env -> Value.t -> Tacexpr.intro_patterns
+ Loc.t -> Environ.env -> Evd.evar_map -> Value.t -> Tacexpr.intro_patterns
-val coerce_to_hyp : Environ.env -> Value.t -> Id.t
+val coerce_to_hyp : Environ.env -> Evd.evar_map -> Value.t -> Id.t
-val coerce_to_hyp_list : Environ.env -> Value.t -> Id.t list
+val coerce_to_hyp_list : Environ.env -> Evd.evar_map -> Value.t -> Id.t list
-val coerce_to_reference : Environ.env -> Value.t -> Globnames.global_reference
+val coerce_to_reference : Environ.env -> Evd.evar_map -> Value.t -> Globnames.global_reference
-val coerce_to_quantified_hypothesis : Value.t -> quantified_hypothesis
+val coerce_to_quantified_hypothesis : Evd.evar_map -> Value.t -> quantified_hypothesis
-val coerce_to_decl_or_quant_hyp : Environ.env -> Value.t -> quantified_hypothesis
+val coerce_to_decl_or_quant_hyp : Environ.env -> Evd.evar_map -> Value.t -> quantified_hypothesis
val coerce_to_int_or_var_list : Value.t -> int or_var list
(** {5 Missing generic arguments} *)
-val wit_constr_context : (Empty.t, Empty.t, constr) genarg_type
+val wit_constr_context : (Empty.t, Empty.t, EConstr.constr) genarg_type
val wit_constr_under_binders : (Empty.t, Empty.t, constr_under_binders) genarg_type
diff --git a/plugins/ltac/tacexpr.mli b/plugins/ltac/tacexpr.mli
index e23992a80..8aefe7605 100644
--- a/plugins/ltac/tacexpr.mli
+++ b/plugins/ltac/tacexpr.mli
@@ -120,9 +120,9 @@ type glob_constr_pattern_and_expr = binding_bound_vars * glob_constr_and_expr *
type 'a delayed_open = 'a Tactypes.delayed_open =
{ delayed : 'r. Environ.env -> 'r Sigma.t -> ('a, 'r) Sigma.sigma }
-type delayed_open_constr_with_bindings = Term.constr with_bindings delayed_open
+type delayed_open_constr_with_bindings = EConstr.constr with_bindings delayed_open
-type delayed_open_constr = Term.constr delayed_open
+type delayed_open_constr = EConstr.constr delayed_open
type intro_pattern = delayed_open_constr intro_pattern_expr located
type intro_patterns = delayed_open_constr intro_pattern_expr located list
@@ -354,7 +354,7 @@ type raw_tactic_arg =
(** Interpreted tactics *)
-type t_trm = Term.constr
+type t_trm = EConstr.constr
type t_pat = constr_pattern
type t_cst = evaluable_global_reference
type t_ref = ltac_constant located
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index fe10f0c31..50f43931e 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -239,12 +239,12 @@ let pr_value env v =
else if has_type v (topwit wit_constr_context) then
let c = out_gen (topwit wit_constr_context) v in
match env with
- | Some (env,sigma) -> pr_lconstr_env env sigma c
+ | Some (env,sigma) -> pr_leconstr_env env sigma c
| _ -> str "a term"
else if has_type v (topwit wit_constr) then
let c = out_gen (topwit wit_constr) v in
match env with
- | Some (env,sigma) -> pr_lconstr_env env sigma c
+ | Some (env,sigma) -> pr_leconstr_env env sigma c
| _ -> str "a term"
else if has_type v (topwit wit_constr_under_binders) then
let c = out_gen (topwit wit_constr_under_binders) v in
@@ -282,7 +282,7 @@ let pr_inspect env expr result =
(* Transforms an id into a constr if possible, or fails with Not_found *)
let constr_of_id env id =
- Term.mkVar (let _ = Environ.lookup_named id env in id)
+ EConstr.mkVar (let _ = Environ.lookup_named id env in id)
(** Generic arguments : table of interpretation functions *)
@@ -385,7 +385,7 @@ let interp_ltac_var coerce ist env locid =
with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time")
let interp_ident ist env sigma id =
- try try_interp_ltac_var (coerce_var_to_ident false env) ist (Some (env,sigma)) (dloc,id)
+ try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (dloc,id)
with Not_found -> id
(* Interprets an optional identifier, bound or fresh *)
@@ -394,11 +394,11 @@ let interp_name ist env sigma = function
| Name id -> Name (interp_ident ist env sigma id)
let interp_intro_pattern_var loc ist env sigma id =
- try try_interp_ltac_var (coerce_to_intro_pattern env) ist (Some (env,sigma)) (loc,id)
+ try try_interp_ltac_var (coerce_to_intro_pattern env sigma) ist (Some (env,sigma)) (loc,id)
with Not_found -> IntroNaming (IntroIdentifier id)
let interp_intro_pattern_naming_var loc ist env sigma id =
- try try_interp_ltac_var (coerce_to_intro_pattern_naming env) ist (Some (env,sigma)) (loc,id)
+ try try_interp_ltac_var (coerce_to_intro_pattern_naming env sigma) ist (Some (env,sigma)) (loc,id)
with Not_found -> IntroIdentifier id
let interp_int ist locid =
@@ -423,14 +423,14 @@ let interp_int_or_var_list ist l =
(* Interprets a bound variable (especially an existing hypothesis) *)
let interp_hyp ist env sigma (loc,id as locid) =
(* Look first in lfun for a value coercible to a variable *)
- try try_interp_ltac_var (coerce_to_hyp env) ist (Some (env,sigma)) locid
+ try try_interp_ltac_var (coerce_to_hyp env sigma) ist (Some (env,sigma)) locid
with Not_found ->
(* Then look if bound in the proof context at calling time *)
if is_variable env id then id
else Loc.raise ~loc (Logic.RefinerError (Logic.NoSuchHyp id))
let interp_hyp_list_as_list ist env sigma (loc,id as x) =
- try coerce_to_hyp_list env (Id.Map.find id ist.lfun)
+ try coerce_to_hyp_list env sigma (Id.Map.find id ist.lfun)
with Not_found | CannotCoerceTo _ -> [interp_hyp ist env sigma x]
let interp_hyp_list ist env sigma l =
@@ -445,7 +445,7 @@ let interp_move_location ist env sigma = function
let interp_reference ist env sigma = function
| ArgArg (_,r) -> r
| ArgVar (loc, id) ->
- try try_interp_ltac_var (coerce_to_reference env) ist (Some (env,sigma)) (loc, id)
+ try try_interp_ltac_var (coerce_to_reference env sigma) ist (Some (env,sigma)) (loc, id)
with Not_found ->
try
VarRef (get_id (Environ.lookup_named id env))
@@ -469,7 +469,7 @@ let interp_evaluable ist env sigma = function
end
| ArgArg (r,None) -> r
| ArgVar (loc, id) ->
- try try_interp_ltac_var (coerce_to_evaluable_ref env) ist (Some (env,sigma)) (loc, id)
+ try try_interp_ltac_var (coerce_to_evaluable_ref env sigma) ist (Some (env,sigma)) (loc, id)
with Not_found ->
try try_interp_evaluable env (loc, id)
with Not_found -> error_global_not_found ~loc (qualid_of_ident id)
@@ -540,7 +540,7 @@ let default_fresh_id = Id.of_string "H"
let interp_fresh_id ist env sigma l =
let extract_ident ist env sigma id =
- try try_interp_ltac_var (coerce_to_ident_not_fresh sigma env)
+ try try_interp_ltac_var (coerce_to_ident_not_fresh env sigma)
ist (Some (env,sigma)) (dloc,id)
with Not_found -> id in
let ids = List.map_filter (function ArgVar (_, id) -> Some id | _ -> None) l in
@@ -561,24 +561,24 @@ let interp_fresh_id ist env sigma l =
Tactics.fresh_id_in_env avoid id env
(* Extract the uconstr list from lfun *)
-let extract_ltac_constr_context ist env =
+let extract_ltac_constr_context ist env sigma =
let open Glob_term in
- let add_uconstr id env v map =
+ let add_uconstr id v map =
try Id.Map.add id (coerce_to_uconstr env v) map
with CannotCoerceTo _ -> map
in
- let add_constr id env v map =
+ let add_constr id v map =
try Id.Map.add id (coerce_to_constr env v) map
with CannotCoerceTo _ -> map
in
- let add_ident id env v map =
- try Id.Map.add id (coerce_var_to_ident false env v) map
+ let add_ident id v map =
+ try Id.Map.add id (coerce_var_to_ident false env sigma v) map
with CannotCoerceTo _ -> map
in
let fold id v {idents;typed;untyped} =
- let idents = add_ident id env v idents in
- let typed = add_constr id env v typed in
- let untyped = add_uconstr id env v untyped in
+ let idents = add_ident id v idents in
+ let typed = add_constr id v typed in
+ let untyped = add_uconstr id v untyped in
{ idents ; typed ; untyped }
in
let empty = { idents = Id.Map.empty ;typed = Id.Map.empty ; untyped = Id.Map.empty } in
@@ -586,11 +586,11 @@ let extract_ltac_constr_context ist env =
(** Significantly simpler than [interp_constr], to interpret an
untyped constr, it suffices to adjoin a closure environment. *)
-let interp_uconstr ist env = function
+let interp_uconstr ist env sigma = function
| (term,None) ->
- { closure = extract_ltac_constr_context ist env ; term }
+ { closure = extract_ltac_constr_context ist env sigma; term }
| (_,Some ce) ->
- let ( {typed ; untyped } as closure) = extract_ltac_constr_context ist env in
+ let ( {typed ; untyped } as closure) = extract_ltac_constr_context ist env sigma in
let ltacvars = {
Constrintern.ltac_vars = Id.(Set.union (Map.domain typed) (Map.domain untyped));
ltac_bound = Id.Map.domain ist.lfun;
@@ -598,7 +598,7 @@ let interp_uconstr ist env = function
{ closure ; term = intern_gen WithoutTypeConstraint ~ltacvars env ce }
let interp_gen kind ist allow_patvar flags env sigma (c,ce) =
- let constrvars = extract_ltac_constr_context ist env in
+ let constrvars = extract_ltac_constr_context ist env sigma in
let vars = {
Pretyping.ltac_constrs = constrvars.typed;
Pretyping.ltac_uconstrs = constrvars.untyped;
@@ -639,7 +639,7 @@ let interp_gen kind ist allow_patvar flags env sigma (c,ce) =
(* spiwack: to avoid unnecessary modifications of tacinterp, as this
function already use effect, I call [run] hoping it doesn't mess
up with any assumption. *)
- Proofview.NonLogical.run (db_constr (curr_debug ist) env c);
+ Proofview.NonLogical.run (db_constr (curr_debug ist) env evd c);
(evd,c)
let constr_flags () = {
@@ -691,7 +691,9 @@ let interp_pure_open_constr ist =
let interp_typed_pattern ist env sigma (_,c,_) =
let sigma, c =
interp_gen WithoutTypeConstraint ist true pure_open_constr_flags env sigma c in
- pattern_of_constr env sigma c
+ (** FIXME: it is necessary to be unsafe here because of the way we handle
+ evars in the pretyper. Sometimes they get solved eagerly. *)
+ pattern_of_constr env sigma (EConstr.Unsafe.to_constr c)
(* Interprets a constr expression *)
let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l =
@@ -733,10 +735,10 @@ let interp_closed_typed_pattern_with_occurrences ist env sigma (occs, a) =
prioritary to an evaluable reference and otherwise to a constr
(it is an encoding to satisfy the "union" type given to Simpl) *)
let coerce_eval_ref_or_constr x =
- try Inl (coerce_to_evaluable_ref env x)
+ try Inl (coerce_to_evaluable_ref env sigma x)
with CannotCoerceTo _ ->
let c = coerce_to_closed_constr env x in
- Inr (pattern_of_constr env sigma c) in
+ Inr (pattern_of_constr env sigma (EConstr.to_constr sigma c)) in
(try try_interp_ltac_var coerce_eval_ref_or_constr ist (Some (env,sigma)) (loc,id)
with Not_found ->
error_global_not_found ~loc (qualid_of_ident id))
@@ -789,9 +791,11 @@ let interp_may_eval f ist env sigma = function
(try
let (sigma,ic) = f ist env sigma c in
let ctxt = coerce_to_constr_context (Id.Map.find s ist.lfun) in
+ let ctxt = EConstr.Unsafe.to_constr ctxt in
let evdref = ref sigma in
+ let ic = EConstr.Unsafe.to_constr ic in
let c = subst_meta [Constr_matching.special_meta,ic] ctxt in
- let c = Typing.e_solve_evars env evdref c in
+ let c = Typing.e_solve_evars env evdref (EConstr.of_constr c) in
!evdref , c
with
| Not_found ->
@@ -799,7 +803,8 @@ let interp_may_eval f ist env sigma = function
(str "Unbound context identifier" ++ pr_id s ++ str"."))
| ConstrTypeOf c ->
let (sigma,c_interp) = f ist env sigma c in
- Typing.type_of ~refresh:true env sigma c_interp
+ let (sigma, t) = Typing.type_of ~refresh:true env sigma c_interp in
+ (sigma, t)
| ConstrTerm c ->
try
f ist env sigma c
@@ -829,7 +834,7 @@ let interp_constr_may_eval ist env sigma c =
(* spiwack: to avoid unnecessary modifications of tacinterp, as this
function already use effect, I call [run] hoping it doesn't mess
up with any assumption. *)
- Proofview.NonLogical.run (db_constr (curr_debug ist) env csr);
+ Proofview.NonLogical.run (db_constr (curr_debug ist) env sigma csr);
sigma , csr
end
@@ -841,10 +846,10 @@ let rec message_of_value v =
Ftactic.return (str "<tactic>")
else if has_type v (topwit wit_constr) then
let v = out_gen (topwit wit_constr) v in
- Ftactic.nf_enter {enter = begin fun gl -> Ftactic.return (pr_constr_env (pf_env gl) (project gl) v) end }
+ Ftactic.enter {enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) v) end }
else if has_type v (topwit wit_constr_under_binders) then
let c = out_gen (topwit wit_constr_under_binders) v in
- Ftactic.nf_enter { enter = begin fun gl ->
+ Ftactic.enter { enter = begin fun gl ->
Ftactic.return (pr_constr_under_binders_env (pf_env gl) (project gl) c)
end }
else if has_type v (topwit wit_unit) then
@@ -853,22 +858,25 @@ let rec message_of_value v =
Ftactic.return (int (out_gen (topwit wit_int) v))
else if has_type v (topwit wit_intro_pattern) then
let p = out_gen (topwit wit_intro_pattern) v in
- let print env sigma c = pr_constr_env env sigma (fst (Tactics.run_delayed env Evd.empty c)) in
- Ftactic.nf_enter { enter = begin fun gl ->
+ let print env sigma c =
+ let (c, sigma) = Tactics.run_delayed env sigma c in
+ pr_econstr_env env sigma c
+ in
+ Ftactic.enter { enter = begin fun gl ->
Ftactic.return (Miscprint.pr_intro_pattern (fun c -> print (pf_env gl) (project gl) c) p)
end }
else if has_type v (topwit wit_constr_context) then
let c = out_gen (topwit wit_constr_context) v in
- Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_constr_env (pf_env gl) (project gl) c) end }
+ Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_econstr_env (pf_env gl) (project gl) c) end }
else if has_type v (topwit wit_uconstr) then
let c = out_gen (topwit wit_uconstr) v in
- Ftactic.nf_enter { enter = begin fun gl ->
+ Ftactic.enter { enter = begin fun gl ->
Ftactic.return (pr_closed_glob_env (pf_env gl)
(project gl) c)
end }
else if has_type v (topwit wit_var) then
let id = out_gen (topwit wit_var) v in
- Ftactic.nf_enter { enter = begin fun gl -> Ftactic.return (pr_id id) end }
+ Ftactic.enter { enter = begin fun gl -> Ftactic.return (pr_id id) end }
else match Value.to_list v with
| Some l ->
Ftactic.List.map message_of_value l >>= fun l ->
@@ -933,7 +941,7 @@ and interp_or_and_intro_pattern ist env sigma = function
and interp_intro_pattern_list_as_list ist env sigma = function
| [loc,IntroNaming (IntroIdentifier id)] as l ->
- (try sigma, coerce_to_intro_pattern_list loc env (Id.Map.find id ist.lfun)
+ (try sigma, coerce_to_intro_pattern_list loc env sigma (Id.Map.find id ist.lfun)
with Not_found | CannotCoerceTo _ ->
List.fold_map (interp_intro_pattern ist env) sigma l)
| l -> List.fold_map (interp_intro_pattern ist env) sigma l
@@ -945,7 +953,7 @@ let interp_intro_pattern_naming_option ist env sigma = function
let interp_or_and_intro_pattern_option ist env sigma = function
| None -> sigma, None
| Some (ArgVar (loc,id)) ->
- (match coerce_to_intro_pattern env (Id.Map.find id ist.lfun) with
+ (match coerce_to_intro_pattern env sigma (Id.Map.find id ist.lfun) with
| IntroAction (IntroOrAndPattern l) -> sigma, Some (loc,l)
| _ ->
user_err ~loc (str "Cannot coerce to a disjunctive/conjunctive pattern."))
@@ -963,31 +971,25 @@ let interp_in_hyp_as ist env sigma (id,ipat) =
let sigma, ipat = interp_intro_pattern_option ist env sigma ipat in
sigma,(interp_hyp ist env sigma id,ipat)
-let interp_quantified_hypothesis ist = function
- | AnonHyp n -> AnonHyp n
- | NamedHyp id ->
- try try_interp_ltac_var coerce_to_quantified_hypothesis ist None(dloc,id)
- with Not_found -> NamedHyp id
-
-let interp_binding_name ist = function
+let interp_binding_name ist sigma = function
| AnonHyp n -> AnonHyp n
| NamedHyp id ->
(* If a name is bound, it has to be a quantified hypothesis *)
(* user has to use other names for variables if these ones clash with *)
(* a name intented to be used as a (non-variable) identifier *)
- try try_interp_ltac_var coerce_to_quantified_hypothesis ist None(dloc,id)
+ try try_interp_ltac_var (coerce_to_quantified_hypothesis sigma) ist None(dloc,id)
with Not_found -> NamedHyp id
let interp_declared_or_quantified_hypothesis ist env sigma = function
| AnonHyp n -> AnonHyp n
| NamedHyp id ->
try try_interp_ltac_var
- (coerce_to_decl_or_quant_hyp env) ist (Some (env,sigma)) (dloc,id)
+ (coerce_to_decl_or_quant_hyp env sigma) ist (Some (env,sigma)) (dloc,id)
with Not_found -> NamedHyp id
let interp_binding ist env sigma (loc,b,c) =
let sigma, c = interp_open_constr ist env sigma c in
- sigma, (loc,interp_binding_name ist b,c)
+ sigma, (loc,interp_binding_name ist sigma b,c)
let interp_bindings ist env sigma = function
| NoBindings ->
@@ -1213,7 +1215,7 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with
tclSHOWHYPS (Proofview.V82.of_tactic (interp_tactic ist tac))
end
| TacAbstract (tac,ido) ->
- Proofview.Goal.nf_enter { enter = begin fun gl -> Tactics.tclABSTRACT
+ Proofview.Goal.enter { enter = begin fun gl -> Tactics.tclABSTRACT
(Option.map (interp_ident ist (pf_env gl) (project gl)) ido) (interp_tactic ist tac)
end }
| TacThen (t1,t) ->
@@ -1355,7 +1357,7 @@ and interp_tacarg ist arg : Val.t Ftactic.t =
Ftactic.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
- let c = interp_uconstr ist env c in
+ let c = interp_uconstr ist env (Sigma.to_evar_map sigma) c in
let Sigma (c, sigma, p) = (type_uconstr ist c).delayed env sigma in
Sigma (Ftactic.return (Value.of_constr c), sigma, p)
end }
@@ -1528,7 +1530,7 @@ and interp_match ist lz constr lmr =
(* Interprets the Match Context expressions *)
and interp_match_goal ist lz lr lmr =
- Ftactic.nf_enter { enter = begin fun gl ->
+ Ftactic.enter { enter = begin fun gl ->
let sigma = project gl in
let env = Proofview.Goal.env gl in
let hyps = Proofview.Goal.hyps gl in
@@ -1593,7 +1595,7 @@ and interp_genarg_var_list ist x =
end }
(* Interprets tactic expressions : returns a "constr" *)
-and interp_ltac_constr ist e : constr Ftactic.t =
+and interp_ltac_constr ist e : EConstr.t Ftactic.t =
let (>>=) = Ftactic.bind in
begin Proofview.tclORELSE
(val_interp ist e)
@@ -1621,7 +1623,7 @@ and interp_ltac_constr ist e : constr Ftactic.t =
debugging_step ist (fun () ->
Pptactic.pr_glob_tactic env e ++ fnl() ++
str " has value " ++ fnl() ++
- pr_constr_env env sigma cresult)
+ pr_econstr_env env sigma cresult)
end <*>
Ftactic.return cresult
with CannotCoerceTo _ ->
@@ -1641,7 +1643,8 @@ and name_atomic ?env tacexpr tac : unit Proofview.tactic =
| Some e -> Proofview.tclUNIT e
| None -> Proofview.tclENV
end >>= fun env ->
- let name () = Pptactic.pr_atomic_tactic env tacexpr in
+ Proofview.tclEVARMAP >>= fun sigma ->
+ let name () = Pptactic.pr_atomic_tactic env sigma tacexpr in
Proofview.Trace.name_tactic name tac
(* Interprets a primitive tactic *)
@@ -1756,8 +1759,7 @@ and interp_atomic ist tac : unit Proofview.tactic =
(Tactics.generalize_gen cl)) sigma
end }
| TacLetTac (na,c,clp,b,eqpat) ->
- Proofview.V82.nf_evar_goals <*>
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = project gl in
let clp = interp_clause ist env sigma clp in
@@ -1794,7 +1796,6 @@ and interp_atomic ist tac : unit Proofview.tactic =
| TacInductionDestruct (isrec,ev,(l,el)) ->
(* spiwack: some unknown part of destruct needs the goal to be
prenormalised. *)
- Proofview.V82.nf_evar_goals <*>
Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = project gl in
@@ -1830,8 +1831,7 @@ and interp_atomic ist tac : unit Proofview.tactic =
| TacChange (None,c,cl) ->
(* spiwack: until the tactic is in the monad *)
Proofview.Trace.name_tactic (fun () -> Pp.str"<change>") begin
- Proofview.V82.nf_evar_goals <*>
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let is_onhyps = match cl.onhyps with
| None | Some [] -> true
| _ -> false
@@ -1860,7 +1860,6 @@ and interp_atomic ist tac : unit Proofview.tactic =
| TacChange (Some op,c,cl) ->
(* spiwack: until the tactic is in the monad *)
Proofview.Trace.name_tactic (fun () -> Pp.str"<change>") begin
- Proofview.V82.nf_evar_goals <*>
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = project gl in
@@ -1905,7 +1904,7 @@ and interp_atomic ist tac : unit Proofview.tactic =
by))
end }
| TacInversion (DepInversion (k,c,ids),hyp) ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = project gl in
let (sigma,c_interp) =
@@ -2053,7 +2052,7 @@ let interp_constr_with_bindings' ist c = Ftactic.return { delayed = fun env sigm
Sigma.Unsafe.of_pair (c, sigma)
}
-let interp_destruction_arg' ist c = Ftactic.nf_enter { enter = begin fun gl ->
+let interp_destruction_arg' ist c = Ftactic.enter { enter = begin fun gl ->
Ftactic.return (interp_destruction_arg ist gl c)
end }
@@ -2087,8 +2086,8 @@ let () =
register_interp0 wit_ltac interp
let () =
- register_interp0 wit_uconstr (fun ist c -> Ftactic.nf_enter { enter = begin fun gl ->
- Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) c)
+ register_interp0 wit_uconstr (fun ist c -> Ftactic.enter { enter = begin fun gl ->
+ Ftactic.return (interp_uconstr ist (Proofview.Goal.env gl) (Tacmach.New.project gl) c)
end })
(***************************************************************************)
@@ -2112,7 +2111,8 @@ let _ =
if Genarg.has_type arg (glbwit wit_tactic) then
let tac = Genarg.out_gen (glbwit wit_tactic) arg in
let tac = interp_tactic ist tac in
- Pfedit.refine_by_tactic env sigma ty tac
+ let (c, sigma) = Pfedit.refine_by_tactic env sigma ty tac in
+ (EConstr.of_constr c, sigma)
else
failwith "not a tactic"
in
diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli
index adbd1d32b..1e5f6bd42 100644
--- a/plugins/ltac/tacinterp.mli
+++ b/plugins/ltac/tacinterp.mli
@@ -9,6 +9,7 @@
open Names
open Tactic_debug
open Term
+open EConstr
open Tacexpr
open Genarg
open Redexpr
@@ -79,7 +80,7 @@ val interp_bindings : interp_sign -> Environ.env -> Evd.evar_map ->
glob_constr_and_expr bindings -> Evd.evar_map * constr bindings
val interp_open_constr_with_bindings : interp_sign -> Environ.env -> Evd.evar_map ->
- glob_constr_and_expr with_bindings -> Evd.evar_map * constr with_bindings
+ glob_constr_and_expr with_bindings -> Evd.evar_map * EConstr.constr with_bindings
(** Initial call for interpretation *)
diff --git a/plugins/ltac/tactic_debug.ml b/plugins/ltac/tactic_debug.ml
index 5cbddc7f6..b2601ad32 100644
--- a/plugins/ltac/tactic_debug.ml
+++ b/plugins/ltac/tactic_debug.ml
@@ -51,7 +51,7 @@ let db_pr_goal gl =
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl gl in
let penv = print_named_context env in
- let pc = print_constr_env env concl in
+ let pc = print_constr_env env (Tacmach.New.project gl) concl in
str" " ++ hv 0 (penv ++ fnl () ++
str "============================" ++ fnl () ++
str" " ++ pc) ++ fnl ()
@@ -223,11 +223,11 @@ let is_debug db =
return (Int.equal skip 0)
(* Prints a constr *)
-let db_constr debug env c =
+let db_constr debug env sigma c =
let open Proofview.NonLogical in
is_debug debug >>= fun db ->
if db then
- msg_tac_debug (str "Evaluated term: " ++ print_constr_env env c)
+ msg_tac_debug (str "Evaluated term: " ++ print_constr_env env sigma c)
else return ()
(* Prints the pattern rule *)
@@ -247,20 +247,20 @@ let hyp_bound = function
| Name id -> str " (bound to " ++ pr_id id ++ str ")"
(* Prints a matched hypothesis *)
-let db_matched_hyp debug env (id,_,c) ido =
+let db_matched_hyp debug env sigma (id,_,c) ido =
let open Proofview.NonLogical in
is_debug debug >>= fun db ->
if db then
msg_tac_debug (str "Hypothesis " ++ pr_id id ++ hyp_bound ido ++
- str " has been matched: " ++ print_constr_env env c)
+ str " has been matched: " ++ print_constr_env env sigma c)
else return ()
(* Prints the matched conclusion *)
-let db_matched_concl debug env c =
+let db_matched_concl debug env sigma c =
let open Proofview.NonLogical in
is_debug debug >>= fun db ->
if db then
- msg_tac_debug (str "Conclusion has been matched: " ++ print_constr_env env c)
+ msg_tac_debug (str "Conclusion has been matched: " ++ print_constr_env env sigma c)
else return ()
(* Prints a success message when the goal has been matched *)
diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli
index 520fb41ef..7745d9b7b 100644
--- a/plugins/ltac/tactic_debug.mli
+++ b/plugins/ltac/tactic_debug.mli
@@ -11,6 +11,7 @@ open Pattern
open Names
open Tacexpr
open Term
+open EConstr
open Evd
(** TODO: Move those definitions somewhere sensible *)
@@ -34,7 +35,7 @@ val debug_prompt :
val db_initialize : unit Proofview.NonLogical.t
(** Prints a constr *)
-val db_constr : debug_info -> env -> constr -> unit Proofview.NonLogical.t
+val db_constr : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t
(** Prints the pattern rule *)
val db_pattern_rule :
@@ -42,10 +43,10 @@ val db_pattern_rule :
(** Prints a matched hypothesis *)
val db_matched_hyp :
- debug_info -> env -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t
+ debug_info -> env -> evar_map -> Id.t * constr option * constr -> Name.t -> unit Proofview.NonLogical.t
(** Prints the matched conclusion *)
-val db_matched_concl : debug_info -> env -> constr -> unit Proofview.NonLogical.t
+val db_matched_concl : debug_info -> env -> evar_map -> constr -> unit Proofview.NonLogical.t
(** Prints a success message when the goal has been matched *)
val db_mc_pattern_success : debug_info -> unit Proofview.NonLogical.t
diff --git a/plugins/ltac/tactic_matching.ml b/plugins/ltac/tactic_matching.ml
index ef45ee47e..5b5cd06cc 100644
--- a/plugins/ltac/tactic_matching.ml
+++ b/plugins/ltac/tactic_matching.ml
@@ -23,8 +23,8 @@ module NamedDecl = Context.Named.Declaration
substitution mapping corresponding to matched hypotheses. *)
type 'a t = {
subst : Constr_matching.bound_ident_map * Pattern.extended_patvar_map ;
- context : Term.constr Id.Map.t;
- terms : Term.constr Id.Map.t;
+ context : EConstr.constr Id.Map.t;
+ terms : EConstr.constr Id.Map.t;
lhs : 'a;
}
@@ -285,7 +285,7 @@ module PatternMatching (E:StaticEnvironment) = struct
let id = NamedDecl.get_id decl in
let refresh = is_local_def decl in
pattern_match_term refresh pat (NamedDecl.get_type decl) () <*>
- put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*>
+ put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*>
return id
(** [hyp_match_type hypname bodypat typepat hyps] matches a single
@@ -297,7 +297,7 @@ module PatternMatching (E:StaticEnvironment) = struct
| LocalDef (id,body,hyp) ->
pattern_match_term false bodypat body () <*>
pattern_match_term true typepat hyp () <*>
- put_terms (id_map_try_add_name hypname (Term.mkVar id) empty_term_subst) <*>
+ put_terms (id_map_try_add_name hypname (EConstr.mkVar id) empty_term_subst) <*>
return id
| LocalAssum (id,hyp) -> fail
diff --git a/plugins/ltac/tactic_matching.mli b/plugins/ltac/tactic_matching.mli
index 090207bcc..300b546f1 100644
--- a/plugins/ltac/tactic_matching.mli
+++ b/plugins/ltac/tactic_matching.mli
@@ -18,8 +18,8 @@
substitution mapping corresponding to matched hypotheses. *)
type 'a t = {
subst : Constr_matching.bound_ident_map * Pattern.extended_patvar_map ;
- context : Term.constr Names.Id.Map.t;
- terms : Term.constr Names.Id.Map.t;
+ context : EConstr.constr Names.Id.Map.t;
+ terms : EConstr.constr Names.Id.Map.t;
lhs : 'a;
}
@@ -31,7 +31,7 @@ type 'a t = {
val match_term :
Environ.env ->
Evd.evar_map ->
- Term.constr ->
+ EConstr.constr ->
(Tacexpr.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list ->
Tacexpr.glob_tactic_expr t Proofview.tactic
@@ -43,7 +43,7 @@ val match_term :
val match_goal:
Environ.env ->
Evd.evar_map ->
- Context.Named.t ->
- Term.constr ->
+ EConstr.named_context ->
+ EConstr.constr ->
(Tacexpr.binding_bound_vars * Pattern.constr_pattern, Tacexpr.glob_tactic_expr) Tacexpr.match_rule list ->
Tacexpr.glob_tactic_expr t Proofview.tactic
diff --git a/plugins/ltac/tauto.ml b/plugins/ltac/tauto.ml
index fb05fd7d0..dc7ee6a23 100644
--- a/plugins/ltac/tauto.ml
+++ b/plugins/ltac/tauto.ml
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Hipattern
open Names
open Pp
@@ -16,6 +17,7 @@ open Tacexpr
open Tacinterp
open Util
open Tacticals.New
+open Proofview.Notations
let tauto_plugin = "tauto"
let () = Mltop.add_known_module tauto_plugin
@@ -111,19 +113,21 @@ let split = Tactics.split_with_bindings false [Misctypes.NoBindings]
(** Test *)
let is_empty _ ist =
- if is_empty_type (assoc_var "X1" ist) then idtac else fail
+ Proofview.tclEVARMAP >>= fun sigma ->
+ if is_empty_type sigma (assoc_var "X1" ist) then idtac else fail
(* Strictly speaking, this exceeds the propositional fragment as it
matches also equality types (and solves them if a reflexivity) *)
let is_unit_or_eq _ ist =
+ Proofview.tclEVARMAP >>= fun sigma ->
let flags = assoc_flags ist in
let test = if flags.strict_unit then is_unit_type else is_unit_or_eq_type in
- if test (assoc_var "X1" ist) then idtac else fail
+ if test sigma (assoc_var "X1" ist) then idtac else fail
-let bugged_is_binary t =
- isApp t &&
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
+let bugged_is_binary sigma t =
+ isApp sigma t &&
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
let (mib,mip) = Global.lookup_inductive ind in
Int.equal mib.Declarations.mind_nparams 2
@@ -132,21 +136,23 @@ let bugged_is_binary t =
(** Dealing with conjunction *)
let is_conj _ ist =
+ Proofview.tclEVARMAP >>= fun sigma ->
let flags = assoc_flags ist in
let ind = assoc_var "X1" ist in
- if (not flags.binary_mode_bugged_detection || bugged_is_binary ind) &&
- is_conjunction
+ if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma ind) &&
+ is_conjunction sigma
~strict:flags.strict_in_hyp_and_ccl
~onlybinary:flags.binary_mode ind
then idtac
else fail
let flatten_contravariant_conj _ ist =
+ Proofview.tclEVARMAP >>= fun sigma ->
let flags = assoc_flags ist in
let typ = assoc_var "X1" ist in
let c = assoc_var "X2" ist in
let hyp = assoc_var "id" ist in
- match match_with_conjunction
+ match match_with_conjunction sigma
~strict:flags.strict_in_contravariant_hyp
~onlybinary:flags.binary_mode typ
with
@@ -154,27 +160,29 @@ let flatten_contravariant_conj _ ist =
let newtyp = List.fold_right mkArrow args c in
let intros = tclMAP (fun _ -> intro) args in
let by = tclTHENLIST [intros; apply hyp; split; assumption] in
- tclTHENLIST [assert_ ~by newtyp; clear (destVar hyp)]
+ tclTHENLIST [assert_ ~by newtyp; clear (destVar sigma hyp)]
| _ -> fail
(** Dealing with disjunction *)
let is_disj _ ist =
+ Proofview.tclEVARMAP >>= fun sigma ->
let flags = assoc_flags ist in
let t = assoc_var "X1" ist in
- if (not flags.binary_mode_bugged_detection || bugged_is_binary t) &&
- is_disjunction
+ if (not flags.binary_mode_bugged_detection || bugged_is_binary sigma t) &&
+ is_disjunction sigma
~strict:flags.strict_in_hyp_and_ccl
~onlybinary:flags.binary_mode t
then idtac
else fail
let flatten_contravariant_disj _ ist =
+ Proofview.tclEVARMAP >>= fun sigma ->
let flags = assoc_flags ist in
let typ = assoc_var "X1" ist in
let c = assoc_var "X2" ist in
let hyp = assoc_var "id" ist in
- match match_with_disjunction
+ match match_with_disjunction sigma
~strict:flags.strict_in_contravariant_hyp
~onlybinary:flags.binary_mode
typ with
@@ -186,7 +194,7 @@ let flatten_contravariant_disj _ ist =
assert_ ~by typ
in
let tacs = List.mapi map args in
- let tac0 = clear (destVar hyp) in
+ let tac0 = clear (destVar sigma hyp) in
tclTHEN (tclTHENLIST tacs) tac0
| _ -> fail
@@ -217,6 +225,7 @@ let apply_nnpp _ ist =
(Proofview.tclUNIT ())
begin fun () -> try
let nnpp = Universes.constr_of_global (Nametab.global_of_path coq_nnpp_path) in
+ let nnpp = EConstr.of_constr nnpp in
apply nnpp
with Not_found -> tclFAIL 0 (Pp.mt ())
end
diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml
index 6051cb3d3..4b87e6e2e 100644
--- a/plugins/micromega/coq_micromega.ml
+++ b/plugins/micromega/coq_micromega.ml
@@ -20,6 +20,8 @@ open Pp
open Mutils
open Goptions
+module Term = EConstr
+
(**
* Debug flag
*)
@@ -330,6 +332,8 @@ struct
open Coqlib
open Term
+ open Constr
+ open EConstr
(**
* Location of the Coq libraries.
@@ -374,6 +378,7 @@ struct
* ZMicromega.v
*)
+ let gen_constant_in_modules s m n = EConstr.of_constr (gen_constant_in_modules s m n)
let init_constant = gen_constant_in_modules "ZMicromega" init_modules
let constant = gen_constant_in_modules "ZMicromega" coq_modules
let bin_constant = gen_constant_in_modules "ZMicromega" bin_module
@@ -599,12 +604,12 @@ struct
(* A simple but useful getter function *)
- let get_left_construct term =
- match Term.kind_of_term term with
- | Term.Construct((_,i),_) -> (i,[| |])
- | Term.App(l,rst) ->
- (match Term.kind_of_term l with
- | Term.Construct((_,i),_) -> (i,rst)
+ let get_left_construct sigma term =
+ match EConstr.kind sigma term with
+ | Constr.Construct((_,i),_) -> (i,[| |])
+ | Constr.App(l,rst) ->
+ (match EConstr.kind sigma l with
+ | Constr.Construct((_,i),_) -> (i,rst)
| _ -> raise ParseError
)
| _ -> raise ParseError
@@ -613,11 +618,11 @@ struct
(* parse/dump/print from numbers up to expressions and formulas *)
- let rec parse_nat term =
- let (i,c) = get_left_construct term in
+ let rec parse_nat sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.O
- | 2 -> Mc.S (parse_nat (c.(0)))
+ | 2 -> Mc.S (parse_nat sigma (c.(0)))
| i -> raise ParseError
let pp_nat o n = Printf.fprintf o "%i" (CoqToCaml.nat n)
@@ -627,11 +632,11 @@ struct
| Mc.O -> Lazy.force coq_O
| Mc.S p -> Term.mkApp(Lazy.force coq_S,[| dump_nat p |])
- let rec parse_positive term =
- let (i,c) = get_left_construct term in
+ let rec parse_positive sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
- | 1 -> Mc.XI (parse_positive c.(0))
- | 2 -> Mc.XO (parse_positive c.(0))
+ | 1 -> Mc.XI (parse_positive sigma c.(0))
+ | 2 -> Mc.XO (parse_positive sigma c.(0))
| 3 -> Mc.XH
| i -> raise ParseError
@@ -661,12 +666,12 @@ struct
let dump_pair t1 t2 dump_t1 dump_t2 (x,y) =
Term.mkApp(Lazy.force coq_pair,[| t1 ; t2 ; dump_t1 x ; dump_t2 y|])
- let parse_z term =
- let (i,c) = get_left_construct term in
+ let parse_z sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.Z0
- | 2 -> Mc.Zpos (parse_positive c.(0))
- | 3 -> Mc.Zneg (parse_positive c.(0))
+ | 2 -> Mc.Zpos (parse_positive sigma c.(0))
+ | 3 -> Mc.Zneg (parse_positive sigma c.(0))
| i -> raise ParseError
let dump_z x =
@@ -686,10 +691,10 @@ struct
Term.mkApp(Lazy.force coq_Qmake,
[| dump_z q.Micromega.qnum ; dump_positive q.Micromega.qden|])
- let parse_q term =
- match Term.kind_of_term term with
- | Term.App(c, args) -> if Constr.equal c (Lazy.force coq_Qmake) then
- {Mc.qnum = parse_z args.(0) ; Mc.qden = parse_positive args.(1) }
+ let parse_q sigma term =
+ match EConstr.kind sigma term with
+ | Constr.App(c, args) -> if EConstr.eq_constr sigma c (Lazy.force coq_Qmake) then
+ {Mc.qnum = parse_z sigma args.(0) ; Mc.qden = parse_positive sigma args.(1) }
else raise ParseError
| _ -> raise ParseError
@@ -719,27 +724,27 @@ struct
| Mc.CInv t -> Term.mkApp(Lazy.force coq_CInv, [| dump_Rcst t |])
| Mc.COpp t -> Term.mkApp(Lazy.force coq_COpp, [| dump_Rcst t |])
- let rec parse_Rcst term =
- let (i,c) = get_left_construct term in
+ let rec parse_Rcst sigma term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> Mc.C0
| 2 -> Mc.C1
- | 3 -> Mc.CQ (parse_q c.(0))
- | 4 -> Mc.CPlus(parse_Rcst c.(0), parse_Rcst c.(1))
- | 5 -> Mc.CMinus(parse_Rcst c.(0), parse_Rcst c.(1))
- | 6 -> Mc.CMult(parse_Rcst c.(0), parse_Rcst c.(1))
- | 7 -> Mc.CInv(parse_Rcst c.(0))
- | 8 -> Mc.COpp(parse_Rcst c.(0))
+ | 3 -> Mc.CQ (parse_q sigma c.(0))
+ | 4 -> Mc.CPlus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 5 -> Mc.CMinus(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 6 -> Mc.CMult(parse_Rcst sigma c.(0), parse_Rcst sigma c.(1))
+ | 7 -> Mc.CInv(parse_Rcst sigma c.(0))
+ | 8 -> Mc.COpp(parse_Rcst sigma c.(0))
| _ -> raise ParseError
- let rec parse_list parse_elt term =
- let (i,c) = get_left_construct term in
+ let rec parse_list sigma parse_elt term =
+ let (i,c) = get_left_construct sigma term in
match i with
| 1 -> []
- | 2 -> parse_elt c.(1) :: parse_list parse_elt c.(2)
+ | 2 -> parse_elt sigma c.(1) :: parse_list sigma parse_elt c.(2)
| i -> raise ParseError
let rec dump_list typ dump_elt l =
@@ -872,9 +877,9 @@ struct
dump_op o ;
dump_expr typ dump_constant e2|])
- let assoc_const x l =
+ let assoc_const sigma x l =
try
- snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l)
+ snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
with
Not_found -> raise ParseError
@@ -898,35 +903,37 @@ struct
let has_typ gl t1 typ =
let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in
- Constr.equal ty typ
+ EConstr.eq_constr (Tacmach.project gl) ty typ
let is_convertible gl t1 t2 =
Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2
let parse_zop gl (op,args) =
- match kind_of_term op with
- | Const (x,_) -> (assoc_const op zop_table, args.(0) , args.(1))
+ let sigma = Tacmach.project gl in
+ match EConstr.kind sigma op with
+ | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1))
| Ind((n,0),_) ->
- if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z)
+ if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_Z)
then (Mc.OpEq, args.(1), args.(2))
else raise ParseError
| _ -> failwith "parse_zop"
let parse_rop gl (op,args) =
- match kind_of_term op with
- | Const (x,_) -> (assoc_const op rop_table, args.(0) , args.(1))
+ let sigma = Tacmach.project gl in
+ match EConstr.kind sigma op with
+ | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1))
| Ind((n,0),_) ->
- if Constr.equal op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R)
+ if EConstr.eq_constr sigma op (Lazy.force coq_Eq) && is_convertible gl args.(0) (Lazy.force coq_R)
then (Mc.OpEq, args.(1), args.(2))
else raise ParseError
| _ -> failwith "parse_zop"
let parse_qop gl (op,args) =
- (assoc_const op qop_table, args.(0) , args.(1))
+ (assoc_const (Tacmach.project gl) op qop_table, args.(0) , args.(1))
- let is_constant t = (* This is an approx *)
- match kind_of_term t with
+ let is_constant sigma t = (* This is an approx *)
+ match EConstr.kind sigma t with
| Construct(i,_) -> true
| _ -> false
@@ -936,9 +943,9 @@ struct
| Power
| Ukn of string
- let assoc_ops x l =
+ let assoc_ops sigma x l =
try
- snd (List.find (fun (x',y) -> Constr.equal x (Lazy.force x')) l)
+ snd (List.find (fun (x',y) -> EConstr.eq_constr sigma x (Lazy.force x')) l)
with
Not_found -> Ukn "Oups"
@@ -950,12 +957,12 @@ struct
struct
type t = constr list
- let compute_rank_add env v =
+ let compute_rank_add env sigma v =
let rec _add env n v =
match env with
| [] -> ([v],n)
| e::l ->
- if eq_constr e v
+ if eq_constr sigma e v
then (env,n)
else
let (env,n) = _add l ( n+1) v in
@@ -963,13 +970,13 @@ struct
let (env, n) = _add env 1 v in
(env, CamlToCoq.positive n)
- let get_rank env v =
+ let get_rank env sigma v =
let rec _get_rank env n =
match env with
| [] -> raise (Invalid_argument "get_rank")
| e::l ->
- if eq_constr e v
+ if eq_constr sigma e v
then n
else _get_rank l (n+1) in
_get_rank env 1
@@ -985,9 +992,9 @@ struct
* This is the big generic function for expression parsers.
*)
- let parse_expr parse_constant parse_exp ops_spec env term =
+ let parse_expr sigma parse_constant parse_exp ops_spec env term =
if debug
- then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.prterm term);
+ then Feedback.msg_debug (Pp.str "parse_expr: " ++ Printer.pr_leconstr term);
(*
let constant_or_variable env term =
@@ -998,7 +1005,7 @@ struct
(Mc.PEX n , env) in
*)
let parse_variable env term =
- let (env,n) = Env.compute_rank_add env term in
+ let (env,n) = Env.compute_rank_add env sigma term in
(Mc.PEX n , env) in
let rec parse_expr env term =
@@ -1009,12 +1016,12 @@ struct
try (Mc.PEc (parse_constant term) , env)
with ParseError ->
- match kind_of_term term with
+ match EConstr.kind sigma term with
| App(t,args) ->
(
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Const c ->
- ( match assoc_ops t ops_spec with
+ ( match assoc_ops sigma t ops_spec with
| Binop f -> combine env f (args.(0),args.(1))
| Opp -> let (expr,env) = parse_expr env args.(0) in
(Mc.PEopp expr, env)
@@ -1026,12 +1033,12 @@ struct
(power , env)
with e when CErrors.noncritical e ->
(* if the exponent is a variable *)
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
+ let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env)
end
| Ukn s ->
if debug
then (Printf.printf "unknown op: %s\n" s; flush stdout;);
- let (env,n) = Env.compute_rank_add env term in (Mc.PEX n, env)
+ let (env,n) = Env.compute_rank_add env sigma term in (Mc.PEX n, env)
)
| _ -> parse_variable env term
)
@@ -1074,60 +1081,60 @@ struct
(* coq_Rdiv , (fun x y -> Mc.CMult(x,Mc.CInv y)) ;*)
]
- let rec rconstant term =
- match Term.kind_of_term term with
+ let rec rconstant sigma term =
+ match EConstr.kind sigma term with
| Const x ->
- if Constr.equal term (Lazy.force coq_R0)
+ if EConstr.eq_constr sigma term (Lazy.force coq_R0)
then Mc.C0
- else if Constr.equal term (Lazy.force coq_R1)
+ else if EConstr.eq_constr sigma term (Lazy.force coq_R1)
then Mc.C1
else raise ParseError
| App(op,args) ->
begin
try
(* the evaluation order is important in the following *)
- let f = assoc_const op rconst_assoc in
- let a = rconstant args.(0) in
- let b = rconstant args.(1) in
+ let f = assoc_const sigma op rconst_assoc in
+ let a = rconstant sigma args.(0) in
+ let b = rconstant sigma args.(1) in
f a b
with
ParseError ->
match op with
- | op when Constr.equal op (Lazy.force coq_Rinv) ->
- let arg = rconstant args.(0) in
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_Rinv) ->
+ let arg = rconstant sigma args.(0) in
if Mc.qeq_bool (Mc.q_of_Rcst arg) {Mc.qnum = Mc.Z0 ; Mc.qden = Mc.XH}
then raise ParseError (* This is a division by zero -- no semantics *)
else Mc.CInv(arg)
- | op when Constr.equal op (Lazy.force coq_IQR) -> Mc.CQ (parse_q args.(0))
- | op when Constr.equal op (Lazy.force coq_IZR) -> Mc.CZ (parse_z args.(0))
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IQR) -> Mc.CQ (parse_q sigma args.(0))
+ | op when EConstr.eq_constr sigma op (Lazy.force coq_IZR) -> Mc.CZ (parse_z sigma args.(0))
| _ -> raise ParseError
end
| _ -> raise ParseError
- let rconstant term =
+ let rconstant sigma term =
if debug
- then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.prterm term ++ fnl ());
- let res = rconstant term in
+ then Feedback.msg_debug (Pp.str "rconstant: " ++ Printer.pr_leconstr term ++ fnl ());
+ let res = rconstant sigma term in
if debug then
(Printf.printf "rconstant -> %a\n" pp_Rcst res ; flush stdout) ;
res
- let parse_zexpr = parse_expr
- zconstant
+ let parse_zexpr sigma = parse_expr sigma
+ (zconstant sigma)
(fun expr x ->
- let exp = (parse_z x) in
+ let exp = (parse_z sigma x) in
match exp with
| Mc.Zneg _ -> Mc.PEc Mc.Z0
| _ -> Mc.PEpow(expr, Mc.Z.to_N exp))
zop_spec
- let parse_qexpr = parse_expr
- qconstant
+ let parse_qexpr sigma = parse_expr sigma
+ (qconstant sigma)
(fun expr x ->
- let exp = parse_z x in
+ let exp = parse_z sigma x in
match exp with
| Mc.Zneg _ ->
begin
@@ -1139,21 +1146,22 @@ struct
Mc.PEpow(expr,exp))
qop_spec
- let parse_rexpr = parse_expr
- rconstant
+ let parse_rexpr sigma = parse_expr sigma
+ (rconstant sigma)
(fun expr x ->
- let exp = Mc.N.of_nat (parse_nat x) in
+ let exp = Mc.N.of_nat (parse_nat sigma x) in
Mc.PEpow(expr,exp))
rop_spec
let parse_arith parse_op parse_expr env cstr gl =
+ let sigma = Tacmach.project gl in
if debug
- then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.prterm cstr ++ fnl ());
- match kind_of_term cstr with
+ then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr cstr ++ fnl ());
+ match EConstr.kind sigma cstr with
| App(op,args) ->
let (op,lhs,rhs) = parse_op gl (op,args) in
- let (e1,env) = parse_expr env lhs in
- let (e2,env) = parse_expr env rhs in
+ let (e1,env) = parse_expr sigma env lhs in
+ let (e2,env) = parse_expr sigma env rhs in
({Mc.flhs = e1; Mc.fop = op;Mc.frhs = e2},env)
| _ -> failwith "error : parse_arith(2)"
@@ -1191,6 +1199,7 @@ struct
*)
let parse_formula gl parse_atom env tg term =
+ let sigma = Tacmach.project gl in
let parse_atom env tg t =
try
@@ -1200,33 +1209,33 @@ struct
let is_prop term =
let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in
- Term.is_prop_sort sort in
+ Sorts.is_prop sort in
let rec xparse_formula env tg term =
- match kind_of_term term with
+ match EConstr.kind sigma term with
| App(l,rst) ->
(match rst with
- | [|a;b|] when eq_constr l (Lazy.force coq_and) ->
+ | [|a;b|] when eq_constr sigma l (Lazy.force coq_and) ->
let f,env,tg = xparse_formula env tg a in
let g,env, tg = xparse_formula env tg b in
mkformula_binary mkC term f g,env,tg
- | [|a;b|] when eq_constr l (Lazy.force coq_or) ->
+ | [|a;b|] when eq_constr sigma l (Lazy.force coq_or) ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkD term f g,env,tg
- | [|a|] when eq_constr l (Lazy.force coq_not) ->
+ | [|a|] when eq_constr sigma l (Lazy.force coq_not) ->
let (f,env,tg) = xparse_formula env tg a in (N(f), env,tg)
- | [|a;b|] when eq_constr l (Lazy.force coq_iff) ->
+ | [|a;b|] when eq_constr sigma l (Lazy.force coq_iff) ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkIff term f g,env,tg
| _ -> parse_atom env tg term)
- | Prod(typ,a,b) when not (Termops.dependent (mkRel 1) b)->
+ | Prod(typ,a,b) when Vars.noccurn sigma 1 b ->
let f,env,tg = xparse_formula env tg a in
let g,env,tg = xparse_formula env tg b in
mkformula_binary mkI term f g,env,tg
- | _ when eq_constr term (Lazy.force coq_True) -> (TT,env,tg)
- | _ when eq_constr term (Lazy.force coq_False) -> (FF,env,tg)
+ | _ when eq_constr sigma term (Lazy.force coq_True) -> (TT,env,tg)
+ | _ when eq_constr sigma term (Lazy.force coq_False) -> (FF,env,tg)
| _ when is_prop term -> X(term),env,tg
| _ -> raise ParseError
in
@@ -1246,10 +1255,10 @@ struct
xdump f
- let prop_env_of_formula form =
+ let prop_env_of_formula sigma form =
let rec doit env = function
| TT | FF | A(_,_,_) -> env
- | X t -> fst (Env.compute_rank_add env t)
+ | X t -> fst (Env.compute_rank_add env sigma t)
| C(f1,f2) | D(f1,f2) | I(f1,_,f2) ->
doit (doit env f1) f2
| N f -> doit env f in
@@ -1380,14 +1389,22 @@ let dump_rexpr = lazy
*)
-let rec make_goal_of_formula dexpr form =
+let prodn n env b =
+ let rec prodrec = function
+ | (0, env, b) -> b
+ | (n, ((v,t)::l), b) -> prodrec (n-1, l, mkProd (v,t,b))
+ | _ -> assert false
+ in
+ prodrec (n,env,b)
+
+let make_goal_of_formula sigma dexpr form =
let vars_idx =
List.mapi (fun i v -> (v, i+1)) (ISet.elements (var_env_of_formula form)) in
(* List.iter (fun (v,i) -> Printf.fprintf stdout "var %i has index %i\n" v i) vars_idx ;*)
- let props = prop_env_of_formula form in
+ let props = prop_env_of_formula sigma form in
let vars_n = List.map (fun (_,i) -> (Names.id_of_string (Printf.sprintf "__x%i" i)) , dexpr.interp_typ) vars_idx in
let props_n = List.mapi (fun i _ -> (Names.id_of_string (Printf.sprintf "__p%i" (i+1))) , Term.mkProp) props in
@@ -1428,7 +1445,7 @@ let rec make_goal_of_formula dexpr form =
| I(x,_,y) -> mkArrow (xdump pi xi x) (xdump (pi+1) (xi+1) y)
| N(x) -> mkArrow (xdump pi xi x) (Lazy.force coq_False)
| A(x,_,_) -> dump_cstr xi x
- | X(t) -> let idx = Env.get_rank props t in
+ | X(t) -> let idx = Env.get_rank props sigma t in
mkRel (pi+idx) in
let nb_vars = List.length vars_n in
@@ -1437,13 +1454,13 @@ let rec make_goal_of_formula dexpr form =
(* Printf.fprintf stdout "NBProps : %i\n" nb_props ;*)
let subst_prop p =
- let idx = Env.get_rank props p in
+ let idx = Env.get_rank props sigma p in
mkVar (Names.id_of_string (Printf.sprintf "__p%i" idx)) in
let form' = map_prop subst_prop form in
- (Term.prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n)
- (Term.prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n)
+ (prodn nb_props (List.map (fun (x,y) -> Names.Name x,y) props_n)
+ (prodn nb_vars (List.map (fun (x,y) -> Names.Name x,y) vars_n)
(xdump (List.length vars_n) 0 form)),
List.rev props_n, List.rev var_name_pos,form')
@@ -1517,19 +1534,19 @@ let rec apply_ids t ids =
| [] -> t
| i::ids -> apply_ids (Term.mkApp(t,[| Term.mkVar i |])) ids
-let coq_Node = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Node")
-let coq_Leaf = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Leaf")
-let coq_Empty = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty")
-
-let coq_VarMap = lazy
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t")
+let coq_Node =
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Node"))
+let coq_Leaf =
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"];["VarMap"]] "Leaf"))
+let coq_Empty =
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ;"VarMap"];["VarMap"]] "Empty"))
+
+let coq_VarMap =
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t"))
let rec dump_varmap typ m =
@@ -1687,7 +1704,8 @@ let rec mk_topo_order le l =
| (Some v,l') -> v :: (mk_topo_order le l')
-let topo_sort_constr l = mk_topo_order Termops.dependent l
+let topo_sort_constr l =
+ mk_topo_order (fun c t -> Termops.dependent Evd.empty (** FIXME *) (EConstr.of_constr c) (EConstr.of_constr t)) l
(**
@@ -1903,7 +1921,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
let formula_typ = (Term.mkApp(Lazy.force coq_Cstr, [|spec.coeff|])) in
let ff = dump_formula formula_typ
(dump_cstr spec.typ spec.dump_coeff) ff in
- Feedback.msg_notice (Printer.prterm ff);
+ Feedback.msg_notice (Printer.pr_leconstr ff);
Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff
end;
@@ -1928,7 +1946,7 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2
let formula_typ = (Term.mkApp( Lazy.force coq_Cstr,[| spec.coeff|])) in
let ff' = dump_formula formula_typ
(dump_cstr spec.typ spec.dump_coeff) ff' in
- Feedback.msg_notice (Printer.prterm ff');
+ Feedback.msg_notice (Printer.pr_leconstr ff');
Printf.fprintf stdout "cnf : %a\n" (pp_cnf (fun o _ -> ())) cnf_ff'
end;
@@ -1962,6 +1980,7 @@ let micromega_gen
spec dumpexpr prover tac =
Proofview.Goal.nf_enter { enter = begin fun gl ->
let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let sigma = Tacmach.project gl in
let concl = Tacmach.pf_concl gl in
let hyps = Tacmach.pf_hyps_types gl in
try
@@ -1973,7 +1992,7 @@ let micromega_gen
match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with
| None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness")
| Some (ids,ff',res') ->
- let (arith_goal,props,vars,ff_arith) = make_goal_of_formula dumpexpr ff' in
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma dumpexpr ff' in
let intro (id,_) = Tactics.introduction id in
let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
@@ -1986,7 +2005,7 @@ let micromega_gen
micromega_order_change spec res'
(Term.mkApp(Lazy.force coq_list, [|spec.proof_typ|])) env' ff_arith ] in
- let goal_props = List.rev (prop_env_of_formula ff') in
+ let goal_props = List.rev (prop_env_of_formula sigma ff') in
let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
@@ -2046,8 +2065,8 @@ let micromega_order_changer cert env ff =
[
("__ff", ff, Term.mkApp(Lazy.force coq_Formula, [|formula_typ |]));
("__varmap", vm, Term.mkApp
- (Coqlib.gen_constant_in_modules "VarMap"
- [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t", [|typ|]));
+ (EConstr.of_constr (Coqlib.gen_constant_in_modules "VarMap"
+ [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t"), [|typ|]));
("__wit", cert, cert_typ)
]
(Tacmach.pf_concl gl)));
@@ -2070,6 +2089,7 @@ let micromega_genr prover tac =
} in
Proofview.Goal.nf_enter { enter = begin fun gl ->
let gl = Tacmach.New.of_old (fun x -> x) gl in
+ let sigma = Tacmach.project gl in
let concl = Tacmach.pf_concl gl in
let hyps = Tacmach.pf_hyps_types gl in
@@ -2088,7 +2108,7 @@ let micromega_genr prover tac =
(List.filter (fun (n,_) -> List.mem n ids) hyps) concl in
let ff' = abstract_wrt_formula ff' ff in
- let (arith_goal,props,vars,ff_arith) = make_goal_of_formula (Lazy.force dump_rexpr) ff' in
+ let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma (Lazy.force dump_rexpr) ff' in
let intro (id,_) = Tactics.introduction id in
let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in
@@ -2100,7 +2120,7 @@ let micromega_genr prover tac =
let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ;
micromega_order_changer res' env' ff_arith ] in
- let goal_props = List.rev (prop_env_of_formula ff') in
+ let goal_props = List.rev (prop_env_of_formula sigma ff') in
let goal_vars = List.map (fun (_,i) -> List.nth env (i-1)) vars in
diff --git a/plugins/nsatz/g_nsatz.ml4 b/plugins/nsatz/g_nsatz.ml4
index 635237d33..759885253 100644
--- a/plugins/nsatz/g_nsatz.ml4
+++ b/plugins/nsatz/g_nsatz.ml4
@@ -14,5 +14,5 @@ open Names
DECLARE PLUGIN "nsatz_plugin"
TACTIC EXTEND nsatz_compute
-| [ "nsatz_compute" constr(lt) ] -> [ Nsatz.nsatz_compute lt ]
+| [ "nsatz_compute" constr(lt) ] -> [ Nsatz.nsatz_compute (EConstr.Unsafe.to_constr lt) ]
END
diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml
index 36bce780b..cc0c4277e 100644
--- a/plugins/nsatz/nsatz.ml
+++ b/plugins/nsatz/nsatz.ml
@@ -625,6 +625,7 @@ let nsatz lpol =
let return_term t =
let a =
mkApp(gen_constant "CC" ["Init";"Logic"] "refl_equal",[|tllp ();t|]) in
+ let a = EConstr.of_constr a in
generalize [a]
let nsatz_compute t =
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 1afc6500b..7780de712 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -18,6 +18,7 @@ open Util
open Names
open Nameops
open Term
+open EConstr
open Tacticals
open Tacmach
open Tactics
@@ -37,7 +38,7 @@ open OmegaSolver
(* Added by JCF, 09/03/98 *)
let elim_id id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
simplest_elim (Tacmach.New.pf_global id gl)
end }
let resolve_id id gl = Proofview.V82.of_tactic (apply (pf_global gl id)) gl
@@ -172,8 +173,8 @@ let tag_hypothesis,tag_of_hyp, hyp_of_tag, clear_tags =
let hide_constr,find_constr,clear_constr_tables,dump_tables =
let l = ref ([]:(constr * (Id.t * Id.t * bool)) list) in
(fun h id eg b -> l := (h,(id,eg,b)):: !l),
- (fun h ->
- try List.assoc_f eq_constr_nounivs h !l with Not_found -> failwith "find_contr"),
+ (fun sigma h ->
+ try List.assoc_f (eq_constr_nounivs sigma) h !l with Not_found -> failwith "find_contr"),
(fun () -> l := []),
(fun () -> !l)
@@ -197,6 +198,7 @@ let coq_modules =
init_modules @arith_modules @ [logic_dir] @ zarith_base_modules
@ [["Coq"; "omega"; "OmegaLemmas"]]
+let gen_constant_in_modules n m s = EConstr.of_constr (gen_constant_in_modules n m s)
let init_constant = gen_constant_in_modules "Omega" init_modules
let constant = gen_constant_in_modules "Omega" coq_modules
@@ -348,11 +350,18 @@ let coq_not_iff = lazy (constant "not_iff")
let coq_not_not = lazy (constant "not_not")
let coq_imp_simp = lazy (constant "imp_simp")
let coq_iff = lazy (constant "iff")
+let coq_not = lazy (init_constant "not")
+let coq_and = lazy (init_constant "and")
+let coq_or = lazy (init_constant "or")
+let coq_eq = lazy (init_constant "eq")
+let coq_ex = lazy (init_constant "ex")
+let coq_False = lazy (init_constant "False")
+let coq_True = lazy (init_constant "True")
(* uses build_coq_and, build_coq_not, build_coq_or, build_coq_ex *)
(* For unfold *)
-let evaluable_ref_of_constr s c = match kind_of_term (Lazy.force c) with
+let evaluable_ref_of_constr s c = match EConstr.kind Evd.empty (Lazy.force c) with
| Const (kn,u) when Tacred.is_evaluable (Global.env()) (EvalConstRef kn) ->
EvalConstRef kn
| _ -> anomaly ~label:"Coq_omega" (Pp.str (s^" is not an evaluable constant"))
@@ -364,21 +373,21 @@ let sp_Zle = lazy (evaluable_ref_of_constr "Z.le" coq_Zle)
let sp_Zgt = lazy (evaluable_ref_of_constr "Z.gt" coq_Zgt)
let sp_Zge = lazy (evaluable_ref_of_constr "Z.ge" coq_Zge)
let sp_Zlt = lazy (evaluable_ref_of_constr "Z.lt" coq_Zlt)
-let sp_not = lazy (evaluable_ref_of_constr "not" (lazy (build_coq_not ())))
+let sp_not = lazy (evaluable_ref_of_constr "not" coq_not)
let mk_var v = mkVar (Id.of_string v)
let mk_plus t1 t2 = mkApp (Lazy.force coq_Zplus, [| t1; t2 |])
let mk_times t1 t2 = mkApp (Lazy.force coq_Zmult, [| t1; t2 |])
let mk_minus t1 t2 = mkApp (Lazy.force coq_Zminus, [| t1;t2 |])
-let mk_eq t1 t2 = mkApp (Universes.constr_of_global (build_coq_eq ()),
+let mk_eq t1 t2 = mkApp (Lazy.force coq_eq,
[| Lazy.force coq_Z; t1; t2 |])
let mk_le t1 t2 = mkApp (Lazy.force coq_Zle, [| t1; t2 |])
let mk_gt t1 t2 = mkApp (Lazy.force coq_Zgt, [| t1; t2 |])
let mk_inv t = mkApp (Lazy.force coq_Zopp, [| t |])
-let mk_and t1 t2 = mkApp (build_coq_and (), [| t1; t2 |])
-let mk_or t1 t2 = mkApp (build_coq_or (), [| t1; t2 |])
-let mk_not t = mkApp (build_coq_not (), [| t |])
-let mk_eq_rel t1 t2 = mkApp (Universes.constr_of_global (build_coq_eq ()),
+let mk_and t1 t2 = mkApp (Lazy.force coq_and, [| t1; t2 |])
+let mk_or t1 t2 = mkApp (Lazy.force coq_or, [| t1; t2 |])
+let mk_not t = mkApp (Lazy.force coq_not, [| t |])
+let mk_eq_rel t1 t2 = mkApp (Lazy.force coq_eq,
[| Lazy.force coq_comparison; t1; t2 |])
let mk_inj t = mkApp (Lazy.force coq_Z_of_nat, [| t |])
@@ -420,22 +429,23 @@ type result =
the term parts that we manipulate, but rather Var's.
Said otherwise: all constr manipulated here are closed *)
-let destructurate_prop t =
- let c, args = decompose_app t in
- match kind_of_term c, args with
- | _, [_;_;_] when is_global (build_coq_eq ()) c -> Kapp (Eq,args)
+let destructurate_prop sigma t =
+ let eq_constr c1 c2 = eq_constr sigma c1 c2 in
+ let c, args = decompose_app sigma t in
+ match EConstr.kind sigma c, args with
+ | _, [_;_;_] when eq_constr (Lazy.force coq_eq) c -> Kapp (Eq,args)
| _, [_;_] when eq_constr c (Lazy.force coq_neq) -> Kapp (Neq,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zne) -> Kapp (Zne,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zle) -> Kapp (Zle,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zlt) -> Kapp (Zlt,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zge) -> Kapp (Zge,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zgt) -> Kapp (Zgt,args)
- | _, [_;_] when eq_constr c (build_coq_and ()) -> Kapp (And,args)
- | _, [_;_] when eq_constr c (build_coq_or ()) -> Kapp (Or,args)
+ | _, [_;_] when eq_constr c (Lazy.force coq_and) -> Kapp (And,args)
+ | _, [_;_] when eq_constr c (Lazy.force coq_or) -> Kapp (Or,args)
| _, [_;_] when eq_constr c (Lazy.force coq_iff) -> Kapp (Iff, args)
- | _, [_] when eq_constr c (build_coq_not ()) -> Kapp (Not,args)
- | _, [] when eq_constr c (build_coq_False ()) -> Kapp (False,args)
- | _, [] when eq_constr c (build_coq_True ()) -> Kapp (True,args)
+ | _, [_] when eq_constr c (Lazy.force coq_not) -> Kapp (Not,args)
+ | _, [] when eq_constr c (Lazy.force coq_False) -> Kapp (False,args)
+ | _, [] when eq_constr c (Lazy.force coq_True) -> Kapp (True,args)
| _, [_;_] when eq_constr c (Lazy.force coq_le) -> Kapp (Le,args)
| _, [_;_] when eq_constr c (Lazy.force coq_lt) -> Kapp (Lt,args)
| _, [_;_] when eq_constr c (Lazy.force coq_ge) -> Kapp (Ge,args)
@@ -451,16 +461,18 @@ let destructurate_prop t =
| Prod (Name _,_,_),[] -> error "Omega: Not a quantifier-free goal"
| _ -> Kufo
-let destructurate_type t =
- let c, args = decompose_app t in
- match kind_of_term c, args with
+let destructurate_type sigma t =
+ let eq_constr c1 c2 = eq_constr sigma c1 c2 in
+ let c, args = decompose_app sigma t in
+ match EConstr.kind sigma c, args with
| _, [] when eq_constr c (Lazy.force coq_Z) -> Kapp (Z,args)
| _, [] when eq_constr c (Lazy.force coq_nat) -> Kapp (Nat,args)
| _ -> Kufo
-let destructurate_term t =
- let c, args = decompose_app t in
- match kind_of_term c, args with
+let destructurate_term sigma t =
+ let eq_constr c1 c2 = eq_constr sigma c1 c2 in
+ let c, args = decompose_app sigma t in
+ match EConstr.kind sigma c, args with
| _, [_;_] when eq_constr c (Lazy.force coq_Zplus) -> Kapp (Zplus,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zmult) -> Kapp (Zmult,args)
| _, [_;_] when eq_constr c (Lazy.force coq_Zminus) -> Kapp (Zminus,args)
@@ -480,15 +492,16 @@ let destructurate_term t =
| Var id,[] -> Kvar id
| _ -> Kufo
-let recognize_number t =
+let recognize_number sigma t =
+ let eq_constr c1 c2 = eq_constr sigma c1 c2 in
let rec loop t =
- match decompose_app t with
+ match decompose_app sigma t with
| f, [t] when eq_constr f (Lazy.force coq_xI) -> one + two * loop t
| f, [t] when eq_constr f (Lazy.force coq_xO) -> two * loop t
| f, [] when eq_constr f (Lazy.force coq_xH) -> one
| _ -> failwith "not a number"
in
- match decompose_app t with
+ match decompose_app sigma t with
| f, [t] when eq_constr f (Lazy.force coq_Zpos) -> loop t
| f, [t] when eq_constr f (Lazy.force coq_Zneg) -> neg (loop t)
| f, [] when eq_constr f (Lazy.force coq_Z0) -> zero
@@ -504,9 +517,9 @@ type constr_path =
| P_ARITY
| P_ARG
-let context operation path (t : constr) =
+let context sigma operation path (t : constr) =
let rec loop i p0 t =
- match (p0,kind_of_term t) with
+ match (p0,EConstr.kind sigma t) with
| (p, Cast (c,k,t)) -> mkCast (loop i p c,k,t)
| ([], _) -> operation i t
| ((P_APP n :: p), App (f,v)) ->
@@ -517,7 +530,7 @@ let context operation path (t : constr) =
let v' = Array.copy v in
v'.(n) <- loop i p v'.(n); (mkCase (ci,q,c,v'))
| ((P_ARITY :: p), App (f,l)) ->
- appvect (loop i p f,l)
+ mkApp (loop i p f,l)
| ((P_ARG :: p), App (f,v)) ->
let v' = Array.copy v in
v'.(0) <- loop i p v'.(0); mkApp (f,v')
@@ -542,8 +555,8 @@ let context operation path (t : constr) =
in
loop 1 path t
-let occurrence path (t : constr) =
- let rec loop p0 t = match (p0,kind_of_term t) with
+let occurrence sigma path (t : constr) =
+ let rec loop p0 t = match (p0,EConstr.kind sigma t) with
| (p, Cast (c,_,_)) -> loop p c
| ([], _) -> t
| ((P_APP n :: p), App (f,v)) -> loop p v.(pred n)
@@ -562,13 +575,13 @@ let occurrence path (t : constr) =
in
loop path t
-let abstract_path typ path t =
+let abstract_path sigma typ path t =
let term_occur = ref (mkRel 0) in
- let abstract = context (fun i t -> term_occur:= t; mkRel i) path t in
+ let abstract = context sigma (fun i t -> term_occur:= t; mkRel i) path t in
mkLambda (Name (Id.of_string "x"), typ, abstract), !term_occur
let focused_simpl path gl =
- let newc = context (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in
+ let newc = context (project gl) (fun i t -> pf_nf gl t) (List.rev path) (pf_concl gl) in
Proofview.V82.of_tactic (convert_concl_no_check newc DEFAULTcast) gl
let focused_simpl path = focused_simpl path
@@ -631,7 +644,7 @@ let mkNewMeta () = mkMeta (Evarutil.new_meta())
let clever_rewrite_base_poly typ p result theorem gl =
let full = pf_concl gl in
- let (abstracted,occ) = abstract_path typ (List.rev p) full in
+ let (abstracted,occ) = abstract_path (project gl) typ (List.rev p) full in
let t =
applist
(mkLambda
@@ -662,8 +675,8 @@ let clever_rewrite_gen_nat p result (t,args) =
let clever_rewrite p vpath t gl =
let full = pf_concl gl in
- let (abstracted,occ) = abstract_path (Lazy.force coq_Z) (List.rev p) full in
- let vargs = List.map (fun p -> occurrence p occ) vpath in
+ let (abstracted,occ) = abstract_path (project gl) (Lazy.force coq_Z) (List.rev p) full in
+ let vargs = List.map (fun p -> occurrence (project gl) p occ) vpath in
let t' = applist(t, (vargs @ [abstracted])) in
exact (applist(t',[mkNewMeta()])) gl
@@ -907,10 +920,10 @@ let rec negate p = function
| Oz i -> [focused_simpl p],Oz(neg i)
| Oufo c -> [], Oufo (mkApp (Lazy.force coq_Zopp, [| c |]))
-let rec transform p t =
+let rec transform sigma p t =
let default isnat t' =
try
- let v,th,_ = find_constr t' in
+ let v,th,_ = find_constr sigma t' in
[clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v
with e when CErrors.noncritical e ->
let v = new_identifier_var ()
@@ -918,29 +931,29 @@ let rec transform p t =
hide_constr t' v th isnat;
[clever_rewrite_base p (mkVar v) (mkVar th)], Oatom v
in
- try match destructurate_term t with
+ try match destructurate_term sigma t with
| Kapp(Zplus,[t1;t2]) ->
- let tac1,t1' = transform (P_APP 1 :: p) t1
- and tac2,t2' = transform (P_APP 2 :: p) t2 in
+ let tac1,t1' = transform sigma (P_APP 1 :: p) t1
+ and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in
let tac,t' = shuffle p (t1',t2') in
tac1 @ tac2 @ tac, t'
| Kapp(Zminus,[t1;t2]) ->
let tac,t =
- transform p
+ transform sigma p
(mkApp (Lazy.force coq_Zplus,
[| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in
Proofview.V82.of_tactic (unfold sp_Zminus) :: tac,t
| Kapp(Zsucc,[t1]) ->
- let tac,t = transform p (mkApp (Lazy.force coq_Zplus,
+ let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus,
[| t1; mk_integer one |])) in
Proofview.V82.of_tactic (unfold sp_Zsucc) :: tac,t
| Kapp(Zpred,[t1]) ->
- let tac,t = transform p (mkApp (Lazy.force coq_Zplus,
+ let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus,
[| t1; mk_integer negone |])) in
Proofview.V82.of_tactic (unfold sp_Zpred) :: tac,t
| Kapp(Zmult,[t1;t2]) ->
- let tac1,t1' = transform (P_APP 1 :: p) t1
- and tac2,t2' = transform (P_APP 2 :: p) t2 in
+ let tac1,t1' = transform sigma (P_APP 1 :: p) t1
+ and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in
begin match t1',t2' with
| (_,Oz n) -> let tac,t' = scalar p n t1' in tac1 @ tac2 @ tac,t'
| (Oz n,_) ->
@@ -951,11 +964,11 @@ let rec transform p t =
| _ -> default false t
end
| Kapp((Zpos|Zneg|Z0),_) ->
- (try ([],Oz(recognize_number t))
+ (try ([],Oz(recognize_number sigma t))
with e when CErrors.noncritical e -> default false t)
| Kvar s -> [],Oatom s
| Kapp(Zopp,[t]) ->
- let tac,t' = transform (P_APP 1 :: p) t in
+ let tac,t' = transform sigma (P_APP 1 :: p) t in
let tac',t'' = negate p t' in
tac @ tac', t''
| Kapp(Z_of_nat,[t']) -> default true t'
@@ -1085,7 +1098,7 @@ let replay_history tactic_normalisation =
let p_initial = [P_APP 2;P_TYPE] in
let tac = shuffle_cancel p_initial e1.body in
let solve_le =
- let not_sup_sup = mkApp (Universes.constr_of_global (build_coq_eq ()),
+ let not_sup_sup = mkApp (Lazy.force coq_eq,
[|
Lazy.force coq_comparison;
Lazy.force coq_Gt;
@@ -1245,7 +1258,7 @@ let replay_history tactic_normalisation =
and eq2 = val_of(decompile orig) in
let vid = unintern_id v in
let theorem =
- mkApp (build_coq_ex (), [|
+ mkApp (Lazy.force coq_ex, [|
Lazy.force coq_Z;
mkLambda
(Name vid,
@@ -1356,15 +1369,15 @@ let replay_history tactic_normalisation =
in
loop
-let normalize p_initial t =
- let (tac,t') = transform p_initial t in
+let normalize sigma p_initial t =
+ let (tac,t') = transform sigma p_initial t in
let (tac',t'') = condense p_initial t' in
let (tac'',t''') = clear_zero p_initial t'' in
tac @ tac' @ tac'' , t'''
-let normalize_equation id flag theorem pos t t1 t2 (tactic,defs) =
+let normalize_equation sigma id flag theorem pos t t1 t2 (tactic,defs) =
let p_initial = [P_APP pos ;P_TYPE] in
- let (tac,t') = normalize p_initial t in
+ let (tac,t') = normalize sigma p_initial t in
let shift_left =
tclTHEN
(Proofview.V82.of_tactic (generalize_tac [mkApp (theorem, [| t1; t2; mkVar id |]) ]))
@@ -1378,35 +1391,39 @@ let normalize_equation id flag theorem pos t t1 t2 (tactic,defs) =
else
(tactic,defs)
+let pf_nf gl c = Tacmach.New.pf_apply Tacred.simpl gl c
+
let destructure_omega gl tac_def (id,c) =
+ let open Tacmach.New in
+ let sigma = project gl in
if String.equal (atompart_of_id id) "State" then
tac_def
else
- try match destructurate_prop c with
+ try match destructurate_prop sigma c with
| Kapp(Eq,[typ;t1;t2])
- when begin match destructurate_type (pf_nf gl typ) with Kapp(Z,[]) -> true | _ -> false end ->
+ when begin match destructurate_type sigma (pf_nf gl typ) with Kapp(Z,[]) -> true | _ -> false end ->
let t = mk_plus t1 (mk_inv t2) in
- normalize_equation
+ normalize_equation sigma
id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def
| Kapp(Zne,[t1;t2]) ->
let t = mk_plus t1 (mk_inv t2) in
- normalize_equation
+ normalize_equation sigma
id DISE (Lazy.force coq_Zne_left) 1 t t1 t2 tac_def
| Kapp(Zle,[t1;t2]) ->
let t = mk_plus t2 (mk_inv t1) in
- normalize_equation
+ normalize_equation sigma
id INEQ (Lazy.force coq_Zle_left) 2 t t1 t2 tac_def
| Kapp(Zlt,[t1;t2]) ->
let t = mk_plus (mk_plus t2 (mk_integer negone)) (mk_inv t1) in
- normalize_equation
+ normalize_equation sigma
id INEQ (Lazy.force coq_Zlt_left) 2 t t1 t2 tac_def
| Kapp(Zge,[t1;t2]) ->
let t = mk_plus t1 (mk_inv t2) in
- normalize_equation
+ normalize_equation sigma
id INEQ (Lazy.force coq_Zge_left) 2 t t1 t2 tac_def
| Kapp(Zgt,[t1;t2]) ->
let t = mk_plus (mk_plus t1 (mk_integer negone)) (mk_inv t2) in
- normalize_equation
+ normalize_equation sigma
id INEQ (Lazy.force coq_Zgt_left) 2 t t1 t2 tac_def
| _ -> tac_def
with e when catchable_exception e -> tac_def
@@ -1419,10 +1436,10 @@ let reintroduce id =
open Proofview.Notations
let coq_omega =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
clear_constr_tables ();
let hyps_types = Tacmach.New.pf_hyps_types gl in
- let destructure_omega = Tacmach.New.of_old destructure_omega gl in
+ let destructure_omega = destructure_omega gl in
let tactic_normalisation, system =
List.fold_left destructure_omega ([],[]) hyps_types in
let prelude,sys =
@@ -1472,10 +1489,11 @@ let coq_omega =
let coq_omega = coq_omega
let nat_inject =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let is_conv = Tacmach.New.pf_apply Reductionops.is_conv gl in
let rec explore p t : unit Proofview.tactic =
- try match destructurate_term t with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ try match destructurate_term sigma t with
| Kapp(Plus,[t1;t2]) ->
Tacticals.New.tclTHENLIST [
Proofview.V82.tactic (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2))
@@ -1511,14 +1529,14 @@ let nat_inject =
]
| Kapp(S,[t']) ->
let rec is_number t =
- try match destructurate_term t with
+ try match destructurate_term sigma t with
Kapp(S,[t]) -> is_number t
| Kapp(O,[]) -> true
| _ -> false
with e when catchable_exception e -> false
in
let rec loop p t : unit Proofview.tactic =
- try match destructurate_term t with
+ try match destructurate_term sigma t with
Kapp(S,[t]) ->
(Tacticals.New.tclTHEN
(Proofview.V82.tactic (clever_rewrite_gen p
@@ -1544,7 +1562,8 @@ let nat_inject =
and loop = function
| [] -> Proofview.tclUNIT ()
| (i,t)::lit ->
- begin try match destructurate_prop t with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ begin try match destructurate_prop sigma t with
Kapp(Le,[t1;t2]) ->
Tacticals.New.tclTHENLIST [
(generalize_tac
@@ -1641,7 +1660,8 @@ let not_binop = function
exception Undecidable
let rec decidability gl t =
- match destructurate_prop t with
+ let open Tacmach.New in
+ match destructurate_prop (project gl) t with
| Kapp(Or,[t1;t2]) ->
mkApp (Lazy.force coq_dec_or, [| t1; t2;
decidability gl t1; decidability gl t2 |])
@@ -1659,7 +1679,7 @@ let rec decidability gl t =
| Kapp(Not,[t1]) ->
mkApp (Lazy.force coq_dec_not, [| t1; decidability gl t1 |])
| Kapp(Eq,[typ;t1;t2]) ->
- begin match destructurate_type (pf_nf gl typ) with
+ begin match destructurate_type (project gl) (pf_nf gl typ) with
| Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |])
| Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |])
| _ -> raise Undecidable
@@ -1671,35 +1691,44 @@ let rec decidability gl t =
| Kapp(True,[]) -> Lazy.force coq_dec_True
| _ -> raise Undecidable
+let fresh_id avoid id gl =
+ fresh_id_in_env avoid id (Proofview.Goal.env gl)
+
let onClearedName id tac =
(* We cannot ensure that hyps can be cleared (because of dependencies), *)
(* so renaming may be necessary *)
Tacticals.New.tclTHEN
(Tacticals.New.tclTRY (clear [id]))
- (Proofview.Goal.nf_enter { enter = begin fun gl ->
- let id = Tacmach.New.of_old (fresh_id [] id) gl in
+ (Proofview.Goal.enter { enter = begin fun gl ->
+ let id = fresh_id [] id gl in
Tacticals.New.tclTHEN (introduction id) (tac id)
end })
let onClearedName2 id tac =
Tacticals.New.tclTHEN
(Tacticals.New.tclTRY (clear [id]))
- (Proofview.Goal.nf_enter { enter = begin fun gl ->
- let id1 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_left")) gl in
- let id2 = Tacmach.New.of_old (fresh_id [] (add_suffix id "_right")) gl in
+ (Proofview.Goal.enter { enter = begin fun gl ->
+ let id1 = fresh_id [] (add_suffix id "_left") gl in
+ let id2 = fresh_id [] (add_suffix id "_right") gl in
Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ]
end })
+let rec is_Prop sigma c = match EConstr.kind sigma c with
+ | Sort s -> Sorts.is_prop (ESorts.kind sigma s)
+ | Cast (c,_,_) -> is_Prop sigma c
+ | _ -> false
+
let destructure_hyps =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
- let decidability = Tacmach.New.of_old decidability gl in
- let pf_nf = Tacmach.New.of_old pf_nf gl in
+ let decidability = decidability gl in
+ let pf_nf = pf_nf gl in
let rec loop = function
| [] -> (Tacticals.New.tclTHEN nat_inject coq_omega)
| decl::lit ->
let i = NamedDecl.get_id decl in
- begin try match destructurate_prop (NamedDecl.get_type decl) with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ begin try match destructurate_prop sigma (NamedDecl.get_type decl) with
| Kapp(False,[]) -> elim_id i
| Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit
| Kapp(Or,[t1;t2]) ->
@@ -1720,7 +1749,7 @@ let destructure_hyps =
| Kimp(t1,t2) ->
(* t1 and t2 might be in Type rather than Prop.
For t1, the decidability check will ensure being Prop. *)
- if is_Prop (type_of t2)
+ if is_Prop sigma (type_of t2)
then
let d1 = decidability t1 in
Tacticals.New.tclTHENLIST [
@@ -1732,7 +1761,7 @@ let destructure_hyps =
else
loop lit
| Kapp(Not,[t]) ->
- begin match destructurate_prop t with
+ begin match destructurate_prop sigma t with
Kapp(Or,[t1;t2]) ->
Tacticals.New.tclTHENLIST [
(generalize_tac
@@ -1789,7 +1818,7 @@ let destructure_hyps =
with Not_found -> loop lit)
| Kapp(Eq,[typ;t1;t2]) ->
if !old_style_flag then begin
- match destructurate_type (pf_nf typ) with
+ match destructurate_type sigma (pf_nf typ) with
| Kapp(Nat,_) ->
Tacticals.New.tclTHENLIST [
(simplest_elim
@@ -1806,7 +1835,7 @@ let destructure_hyps =
]
| _ -> loop lit
end else begin
- match destructurate_type (pf_nf typ) with
+ match destructurate_type sigma (pf_nf typ) with
| Kapp(Nat,_) ->
(Tacticals.New.tclTHEN
(convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|]))
@@ -1832,11 +1861,14 @@ let destructure_hyps =
end }
let destructure_goal =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
- let decidability = Tacmach.New.of_old decidability gl in
+ let decidability = decidability gl in
let rec loop t =
- match destructurate_prop t with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ let prop () = Proofview.tclUNIT (destructurate_prop sigma t) in
+ Proofview.V82.wrap_exceptions prop >>= fun prop ->
+ match prop with
| Kapp(Not,[t]) ->
(Tacticals.New.tclTHEN
(Tacticals.New.tclTHEN (unfold sp_not) intro)
@@ -1851,7 +1883,8 @@ let destructure_goal =
(Proofview.V82.tactic (Tacmach.refine
(mkApp (Lazy.force coq_dec_not_not, [| t; dec; mkNewMeta () |]))))
intro
- with Undecidable -> Tactics.elim_type (build_coq_False ())
+ with Undecidable -> Tactics.elim_type (Lazy.force coq_False)
+ | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
in
Tacticals.New.tclTHEN goal_tac destructure_hyps
in
diff --git a/plugins/quote/g_quote.ml4 b/plugins/quote/g_quote.ml4
index f2c021f59..6c3e66112 100644
--- a/plugins/quote/g_quote.ml4
+++ b/plugins/quote/g_quote.ml4
@@ -23,7 +23,7 @@ let loc = Loc.ghost
let cont = Id.of_string "cont"
let x = Id.of_string "x"
-let make_cont (k : Val.t) (c : Constr.t) =
+let make_cont (k : Val.t) (c : EConstr.t) =
let c = Tacinterp.Value.of_constr c in
let tac = TacCall (loc, ArgVar (loc, cont), [Reference (ArgVar (loc, x))]) in
let ist = { lfun = Id.Map.add cont k (Id.Map.singleton x c); extra = TacStore.empty; } in
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index 6405c8ceb..fc9d70ae7 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -105,6 +105,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Pattern
open Patternops
open Constr_matching
@@ -116,7 +117,8 @@ open Proofview.Notations
We do that lazily, because this code can be linked before
the constants are loaded in the environment *)
-let constant dir s = Coqlib.gen_constant "Quote" ("quote"::dir) s
+let constant dir s =
+ EConstr.of_constr (Coqlib.gen_constant "Quote" ("quote"::dir) s)
let coq_Empty_vm = lazy (constant ["Quote"] "Empty_vm")
let coq_Node_vm = lazy (constant ["Quote"] "Node_vm")
@@ -165,7 +167,7 @@ exchange ?1 and ?2 in the example above)
module ConstrSet = Set.Make(
struct
- type t = constr
+ type t = Constr.constr
let compare = constr_ord
end)
@@ -183,7 +185,7 @@ type inversion_scheme = {
let i_can't_do_that () = error "Quote: not a simple fixpoint"
-let decomp_term c = kind_of_term (Termops.strip_outer_cast c)
+let decomp_term sigma c = EConstr.kind sigma (Termops.strip_outer_cast sigma c)
(*s [compute_lhs typ i nargsi] builds the term \texttt{(C ?nargsi ...
?2 ?1)}, where \texttt{C} is the [i]-th constructor of inductive
@@ -195,8 +197,8 @@ let coerce_meta_out id =
let coerce_meta_in n =
Id.of_string ("M" ^ string_of_int n)
-let compute_lhs typ i nargsi =
- match kind_of_term typ with
+let compute_lhs sigma typ i nargsi =
+ match EConstr.kind sigma typ with
| Ind((sp,0),u) ->
let argsi = Array.init nargsi (fun j -> mkMeta (nargsi - j)) in
mkApp (mkConstructU (((sp,0),i+1),u), argsi)
@@ -205,60 +207,62 @@ let compute_lhs typ i nargsi =
(*s This function builds the pattern from the RHS. Recursive calls are
replaced by meta-variables ?i corresponding to those in the LHS *)
-let compute_rhs bodyi index_of_f =
+let compute_rhs env sigma bodyi index_of_f =
let rec aux c =
- match kind_of_term c with
- | App (j, args) when isRel j && Int.equal (destRel j) index_of_f (* recursive call *) ->
- let i = destRel (Array.last args) in
+ match EConstr.kind sigma c with
+ | App (j, args) when isRel sigma j && Int.equal (destRel sigma j) index_of_f (* recursive call *) ->
+ let i = destRel sigma (Array.last args) in
PMeta (Some (coerce_meta_in i))
| App (f,args) ->
- PApp (pattern_of_constr (Global.env()) Evd.empty f, Array.map aux args)
+ PApp (pattern_of_constr env sigma (EConstr.to_constr sigma f), Array.map aux args)
| Cast (c,_,_) -> aux c
- | _ -> pattern_of_constr (Global.env())(*FIXME*) Evd.empty c
+ | _ -> pattern_of_constr env sigma (EConstr.to_constr sigma c)
in
aux bodyi
(*s Now the function [compute_ivs] itself *)
let compute_ivs f cs gl =
- let cst = try destConst f with DestKO -> i_can't_do_that () in
- let body = Environ.constant_value_in (Global.env()) cst in
- match decomp_term body with
+ let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
+ let (cst, u) = try destConst sigma f with DestKO -> i_can't_do_that () in
+ let u = EInstance.kind sigma u in
+ let body = Environ.constant_value_in (Global.env()) (cst, u) in
+ let body = EConstr.of_constr body in
+ match decomp_term sigma body with
| Fix(([| len |], 0), ([| name |], [| typ |], [| body2 |])) ->
- let (args3, body3) = decompose_lam body2 in
+ let (args3, body3) = decompose_lam sigma body2 in
let nargs3 = List.length args3 in
- let env = Proofview.Goal.env gl in
- let sigma = Tacmach.New.project gl in
let is_conv = Reductionops.is_conv env sigma in
- begin match decomp_term body3 with
+ begin match decomp_term sigma body3 with
| Case(_,p,c,lci) -> (* <p> Case c of c1 ... cn end *)
let n_lhs_rhs = ref []
and v_lhs = ref (None : constr option)
and c_lhs = ref (None : constr option) in
Array.iteri
(fun i ci ->
- let argsi, bodyi = decompose_lam ci in
+ let argsi, bodyi = decompose_lam sigma ci in
let nargsi = List.length argsi in
(* REL (narg3 + nargsi + 1) is f *)
(* REL nargsi+1 to REL nargsi + nargs3 are arguments of f *)
(* REL 1 to REL nargsi are argsi (reverse order) *)
(* First we test if the RHS is the RHS for constants *)
- if isRel bodyi && Int.equal (destRel bodyi) 1 then
- c_lhs := Some (compute_lhs (snd (List.hd args3))
+ if isRel sigma bodyi && Int.equal (destRel sigma bodyi) 1 then
+ c_lhs := Some (compute_lhs sigma (snd (List.hd args3))
i nargsi)
(* Then we test if the RHS is the RHS for variables *)
- else begin match decompose_app bodyi with
+ else begin match decompose_app sigma bodyi with
| vmf, [_; _; a3; a4 ]
- when isRel a3 && isRel a4 && is_conv vmf
- (Lazy.force coq_varmap_find)->
- v_lhs := Some (compute_lhs
+ when isRel sigma a3 && isRel sigma a4 && is_conv vmf
+ (Lazy.force coq_varmap_find) ->
+ v_lhs := Some (compute_lhs sigma
(snd (List.hd args3))
i nargsi)
(* Third case: this is a normal LHS-RHS *)
| _ ->
n_lhs_rhs :=
- (compute_lhs (snd (List.hd args3)) i nargsi,
- compute_rhs bodyi (nargs3 + nargsi + 1))
+ (compute_lhs sigma (snd (List.hd args3)) i nargsi,
+ compute_rhs env sigma bodyi (nargs3 + nargsi + 1))
:: !n_lhs_rhs
end)
lci;
@@ -266,7 +270,7 @@ let compute_ivs f cs gl =
if Option.is_empty !c_lhs && Option.is_empty !v_lhs then i_can't_do_that ();
(* The Cases predicate is a lambda; we assume no dependency *)
- let p = match kind_of_term p with
+ let p = match EConstr.kind sigma p with
| Lambda (_,_,p) -> Termops.pop p
| _ -> p
in
@@ -297,11 +301,11 @@ binary search trees (see file \texttt{Quote.v}) *)
(* First the function to distinghish between constants (closed terms)
and variables (open terms) *)
-let rec closed_under cset t =
- (ConstrSet.mem t cset) ||
- (match (kind_of_term t) with
- | Cast(c,_,_) -> closed_under cset c
- | App(f,l) -> closed_under cset f && Array.for_all (closed_under cset) l
+let rec closed_under sigma cset t =
+ (ConstrSet.mem (EConstr.Unsafe.to_constr t) cset) ||
+ (match EConstr.kind sigma t with
+ | Cast(c,_,_) -> closed_under sigma cset c
+ | App(f,l) -> closed_under sigma cset f && Array.for_all (closed_under sigma cset) l
| _ -> false)
(*s [btree_of_array [| c1; c2; c3; c4; c5 |]] builds the complete
@@ -361,7 +365,7 @@ let path_of_int n =
let rec subterm gl (t : constr) (t' : constr) =
(pf_conv_x gl t t') ||
- (match (kind_of_term t) with
+ (match EConstr.kind (project gl) t with
| App (f,args) -> Array.exists (fun t -> subterm gl t t') args
| Cast(t,_,_) -> (subterm gl t t')
| _ -> false)
@@ -370,9 +374,10 @@ let rec subterm gl (t : constr) (t' : constr) =
(* Since it's a partial order the algoritm of Sort.list won't work !! *)
let rec sort_subterm gl l =
+ let sigma = project gl in
let rec insert c = function
| [] -> [c]
- | (h::t as l) when eq_constr c h -> l (* Avoid doing the same work twice *)
+ | (h::t as l) when EConstr.eq_constr sigma c h -> l (* Avoid doing the same work twice *)
| h::t -> if subterm gl c h then c::h::t else h::(insert c t)
in
match l with
@@ -380,11 +385,15 @@ let rec sort_subterm gl l =
| h::t -> insert h (sort_subterm gl t)
module Constrhash = Hashtbl.Make
- (struct type t = constr
- let equal = eq_constr
- let hash = hash_constr
+ (struct type t = Constr.constr
+ let equal = Term.eq_constr
+ let hash = Term.hash_constr
end)
+let subst_meta subst c =
+ let subst = List.map (fun (i, c) -> i, EConstr.Unsafe.to_constr c) subst in
+ EConstr.of_constr (Termops.subst_meta subst (EConstr.Unsafe.to_constr c))
+
(*s Now we are able to do the inversion itself.
We destructurate the term and use an imperative hashtable
to store leafs that are already encountered.
@@ -392,7 +401,7 @@ module Constrhash = Hashtbl.Make
[ivs : inversion_scheme]\\
[lc: constr list]\\
[gl: goal sigma]\\ *)
-let quote_terms ivs lc =
+let quote_terms env sigma ivs lc =
Coqlib.check_required_library ["Coq";"quote";"Quote"];
let varhash = (Constrhash.create 17 : constr Constrhash.t) in
let varlist = ref ([] : constr list) in (* list of variables *)
@@ -402,34 +411,34 @@ let quote_terms ivs lc =
match l with
| (lhs, rhs)::tail ->
begin try
- let s1 = Id.Map.bindings (matches (Global.env ()) Evd.empty rhs c) in
+ let s1 = Id.Map.bindings (matches env sigma rhs c) in
let s2 = List.map (fun (i,c_i) -> (coerce_meta_out i,aux c_i)) s1
in
- Termops.subst_meta s2 lhs
+ subst_meta s2 lhs
with PatternMatchingFailure -> auxl tail
end
| [] ->
begin match ivs.variable_lhs with
| None ->
begin match ivs.constant_lhs with
- | Some c_lhs -> Termops.subst_meta [1, c] c_lhs
+ | Some c_lhs -> subst_meta [1, c] c_lhs
| None -> anomaly (Pp.str "invalid inversion scheme for quote")
end
| Some var_lhs ->
begin match ivs.constant_lhs with
- | Some c_lhs when closed_under ivs.constants c ->
- Termops.subst_meta [1, c] c_lhs
+ | Some c_lhs when closed_under sigma ivs.constants c ->
+ subst_meta [1, c] c_lhs
| _ ->
begin
- try Constrhash.find varhash c
+ try Constrhash.find varhash (EConstr.Unsafe.to_constr c)
with Not_found ->
let newvar =
- Termops.subst_meta [1, (path_of_int !counter)]
+ subst_meta [1, (path_of_int !counter)]
var_lhs in
begin
incr counter;
varlist := c :: !varlist;
- Constrhash.add varhash c newvar;
+ Constrhash.add varhash (EConstr.Unsafe.to_constr c) newvar;
newvar
end
end
@@ -448,11 +457,13 @@ let quote_terms ivs lc =
let quote f lid =
Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let f = Tacmach.New.pf_global f gl in
- let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in
+ let cl = List.map (fun id -> EConstr.to_constr sigma (Tacmach.New.pf_global id gl)) lid in
let ivs = compute_ivs f cl gl in
let concl = Proofview.Goal.concl gl in
- let quoted_terms = quote_terms ivs [concl] in
+ let quoted_terms = quote_terms env sigma ivs [concl] in
let (p, vm) = match quoted_terms with
| [p], vm -> (p,vm)
| _ -> assert false
@@ -464,10 +475,12 @@ let quote f lid =
let gen_quote cont c f lid =
Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let f = Tacmach.New.pf_global f gl in
- let cl = List.map (fun id -> Tacmach.New.pf_global id gl) lid in
+ let cl = List.map (fun id -> EConstr.to_constr sigma (Tacmach.New.pf_global id gl)) lid in
let ivs = compute_ivs f cl gl in
- let quoted_terms = quote_terms ivs [c] in
+ let quoted_terms = quote_terms env sigma ivs [c] in
let (p, vm) = match quoted_terms with
| [p], vm -> (p,vm)
| _ -> assert false
diff --git a/plugins/romega/const_omega.ml b/plugins/romega/const_omega.ml
index 4935fe4bb..5c68078d7 100644
--- a/plugins/romega/const_omega.ml
+++ b/plugins/romega/const_omega.ml
@@ -353,7 +353,7 @@ let parse_term t =
let parse_rel gl t =
try match destructurate t with
| Kapp("eq",[typ;t1;t2])
- when destructurate (Tacmach.pf_nf gl typ) = Kapp("Z",[]) -> Req (t1,t2)
+ when destructurate (EConstr.Unsafe.to_constr (Tacmach.pf_nf gl (EConstr.of_constr typ))) = Kapp("Z",[]) -> Req (t1,t2)
| Kapp("Zne",[t1;t2]) -> Rne (t1,t2)
| Kapp("Z.le",[t1;t2]) -> Rle (t1,t2)
| Kapp("Z.lt",[t1;t2]) -> Rlt (t1,t2)
diff --git a/plugins/romega/refl_omega.ml b/plugins/romega/refl_omega.ml
index ba882e39a..cfe14b230 100644
--- a/plugins/romega/refl_omega.ml
+++ b/plugins/romega/refl_omega.ml
@@ -740,6 +740,7 @@ and oproposition_of_constr env ((negated,depends,origin,path) as ctxt) gl c =
let reify_gl env gl =
let concl = Tacmach.pf_concl gl in
+ let concl = EConstr.Unsafe.to_constr concl in
let t_concl =
Pnot (oproposition_of_constr env (true,[],id_concl,[O_mono]) gl concl) in
if !debug then begin
@@ -748,6 +749,7 @@ let reify_gl env gl =
end;
let rec loop = function
(i,t) :: lhyps ->
+ let t = EConstr.Unsafe.to_constr t in
let t' = oproposition_of_constr env (false,[],i,[]) gl t in
if !debug then begin
Printf.printf " %s: " (Names.Id.to_string i);
@@ -1222,7 +1224,7 @@ let resolution env full_reified_goal systems_list =
(* variables a introduire *)
let to_introduce = add_stated_equations env solution_tree in
let stated_vars = List.map (fun (v,_,_,_) -> v) to_introduce in
- let l_generalize_arg = List.map (fun (_,t,_,_) -> t) to_introduce in
+ let l_generalize_arg = List.map (fun (_,t,_,_) -> EConstr.of_constr t) to_introduce in
let hyp_stated_vars = List.map (fun (_,_,_,id) -> CCEqua id) to_introduce in
(* L'environnement de base se construit en deux morceaux :
- les variables des équations utiles (et de la conclusion)
@@ -1258,6 +1260,7 @@ let resolution env full_reified_goal systems_list =
let reified =
app coq_interp_sequent
[| reified_concl;env_props_reified;env_terms_reified;reified_goal|] in
+ let reified = EConstr.of_constr reified in
let normalize_equation e =
let rec loop = function
[] -> app (if e.e_negated then coq_p_invert else coq_p_step)
@@ -1281,9 +1284,9 @@ let resolution env full_reified_goal systems_list =
let decompose_tactic = decompose_tree env context solution_tree in
Proofview.V82.of_tactic (Tactics.generalize
- (l_generalize_arg @ List.map Term.mkVar (List.tl l_hyps))) >>
+ (l_generalize_arg @ List.map EConstr.mkVar (List.tl l_hyps))) >>
Proofview.V82.of_tactic (Tactics.change_concl reified) >>
- Proofview.V82.of_tactic (Tactics.apply (app coq_do_omega [|decompose_tactic; normalization_trace|])) >>
+ Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (app coq_do_omega [|decompose_tactic; normalization_trace|]))) >>
show_goal >>
Proofview.V82.of_tactic (Tactics.normalise_vm_in_concl) >>
(*i Alternatives to the previous line:
@@ -1292,7 +1295,7 @@ let resolution env full_reified_goal systems_list =
- Skip the conversion check and rely directly on the QED:
Tacmach.convert_concl_no_check (Lazy.force coq_True) Term.VMcast >>
i*)
- Proofview.V82.of_tactic (Tactics.apply (Lazy.force coq_I))
+ Proofview.V82.of_tactic (Tactics.apply (EConstr.of_constr (Lazy.force coq_I)))
let total_reflexive_omega_tactic gl =
Coqlib.check_required_library ["Coq";"romega";"ROmega"];
diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml
index 35d6768c1..f30f8a943 100644
--- a/plugins/rtauto/refl_tauto.ml
+++ b/plugins/rtauto/refl_tauto.ml
@@ -67,19 +67,18 @@ let l_E_Or = lazy (constant "E_Or")
let l_D_Or = lazy (constant "D_Or")
-let special_whd gl=
- let infos=CClosure.create_clos_infos CClosure.all (pf_env gl) in
- (fun t -> CClosure.whd_val infos (CClosure.inject t))
+let special_whd gl c =
+ Reductionops.clos_whd_flags CClosure.all (pf_env gl) (Tacmach.project gl) c
-let special_nf gl=
- let infos=CClosure.create_clos_infos CClosure.betaiotazeta (pf_env gl) in
- (fun t -> CClosure.norm_val infos (CClosure.inject t))
+let special_nf gl c =
+ Reductionops.clos_norm_flags CClosure.betaiotazeta (pf_env gl) (Tacmach.project gl) c
type atom_env=
{mutable next:int;
mutable env:(constr*int) list}
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
@@ -91,13 +90,17 @@ let make_atom atom_env term=
Atom i
let rec make_form atom_env gls term =
+ let open EConstr in
+ let open Vars in
let normalize=special_nf gls in
let cciterm=special_whd gls term in
- match kind_of_term cciterm with
+ let sigma = Tacmach.project gls in
+ let inj = EConstr.Unsafe.to_constr in
+ match EConstr.kind sigma cciterm with
Prod(_,a,b) ->
- if not (Termops.dependent (mkRel 1) b) &&
+ if noccurn sigma 1 b &&
Retyping.get_sort_family_of
- (pf_env gls) (Tacmach.project gls) a == InProp
+ (pf_env gls) sigma a == InProp
then
let fa=make_form atom_env gls a in
let fb=make_form atom_env gls b in
@@ -114,7 +117,7 @@ let rec make_form atom_env gls term =
| App(hd,argv) when Int.equal (Array.length argv) 2 ->
begin
try
- let ind, _ = destInd hd in
+ let ind, _ = destInd sigma hd in
if Names.eq_ind ind (fst (Lazy.force li_and)) then
let fa=make_form atom_env gls argv.(0) in
let fb=make_form atom_env gls argv.(1) in
@@ -135,7 +138,7 @@ let rec make_hyps atom_env gls lenv = function
| LocalAssum (id,typ)::rest ->
let hrec=
make_hyps atom_env gls (typ::lenv) rest in
- if List.exists (Termops.dependent (mkVar id)) lenv ||
+ if List.exists (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id c) lenv ||
(Retyping.get_sort_family_of
(pf_env gls) (Tacmach.project gls) typ != InProp)
then
@@ -313,6 +316,7 @@ let rtauto_tac gls=
str "Giving proof term to Coq ... ")
end in
let tac_start_time = System.get_time () in
+ let term = EConstr.of_constr term in
let result=
if !check then
Proofview.V82.of_tactic (Tactics.exact_check term) gls
diff --git a/plugins/rtauto/refl_tauto.mli b/plugins/rtauto/refl_tauto.mli
index 9a14ac6c7..092552364 100644
--- a/plugins/rtauto/refl_tauto.mli
+++ b/plugins/rtauto/refl_tauto.mli
@@ -12,13 +12,13 @@ type atom_env=
mutable env:(Term.constr*int) list}
val make_form : atom_env ->
- Proof_type.goal Tacmach.sigma -> Term.types -> Proof_search.form
+ Proof_type.goal Tacmach.sigma -> EConstr.types -> Proof_search.form
val make_hyps :
atom_env ->
Proof_type.goal Tacmach.sigma ->
- Term.types list ->
- Context.Named.t ->
+ EConstr.types list ->
+ EConstr.named_context ->
(Names.Id.t * Proof_search.form) list
val rtauto_tac : Proof_type.tactic
diff --git a/plugins/setoid_ring/g_newring.ml4 b/plugins/setoid_ring/g_newring.ml4
index 707ff79a6..05ab8ab32 100644
--- a/plugins/setoid_ring/g_newring.ml4
+++ b/plugins/setoid_ring/g_newring.ml4
@@ -78,9 +78,7 @@ END
VERNAC COMMAND EXTEND AddSetoidRing CLASSIFIED AS SIDEFF
| [ "Add" "Ring" ident(id) ":" constr(t) ring_mods_opt(l) ] ->
- [ let l = match l with None -> [] | Some l -> l in
- let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in
- add_theory id (ic t) set k cst (pre,post) power sign div]
+ [ let l = match l with None -> [] | Some l -> l in add_theory id t l]
| [ "Print" "Rings" ] => [Vernac_classifier.classify_as_query] -> [
Feedback.msg_notice (strbrk "The following ring structures have been declared:");
Spmap.iter (fun fn fi ->
@@ -93,7 +91,7 @@ END
TACTIC EXTEND ring_lookup
| [ "ring_lookup" tactic0(f) "[" constr_list(lH) "]" ne_constr_list(lrt) ] ->
- [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t]
+ [ let (t,lr) = List.sep_last lrt in ring_lookup f lH lr t ]
END
let pr_field_mod = function
@@ -115,9 +113,7 @@ END
VERNAC COMMAND EXTEND AddSetoidField CLASSIFIED AS SIDEFF
| [ "Add" "Field" ident(id) ":" constr(t) field_mods_opt(l) ] ->
- [ let l = match l with None -> [] | Some l -> l in
- let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods l in
- add_field_theory id (ic t) set k cst_tac inj (pre,post) power sign div]
+ [ let l = match l with None -> [] | Some l -> l in add_field_theory id t l ]
| [ "Print" "Fields" ] => [Vernac_classifier.classify_as_query] -> [
Feedback.msg_notice (strbrk "The following field structures have been declared:");
Spmap.iter (fun fn fi ->
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 87ee66660..dd68eac24 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -12,6 +12,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Vars
open CClosure
open Environ
@@ -43,9 +44,9 @@ let tag_arg tag_rec map subs i c =
| Prot -> mk_atom c
| Rec -> if Int.equal i (-1) then mk_clos subs c else tag_rec c
-let global_head_of_constr c =
- let f, args = decompose_app c in
- try global_of_constr f
+let global_head_of_constr sigma c =
+ let f, args = decompose_app sigma c in
+ try fst (Termops.global_of_constr sigma f)
with Not_found -> anomaly (str "global_head_of_constr")
let global_of_constr_nofail c =
@@ -53,6 +54,7 @@ let global_of_constr_nofail c =
with Not_found -> VarRef (Id.of_string "dummy")
let rec mk_clos_but f_map subs t =
+ let open Term in
match f_map (global_of_constr_nofail t) with
| Some map -> tag_arg (mk_clos_but f_map subs) map subs (-1) t
| None ->
@@ -62,6 +64,7 @@ let rec 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
if n >= Array.length args then mk_atom(mkApp(f, args))
else
let fargs, args' = Array.chop n args in
@@ -82,9 +85,11 @@ let lookup_map map =
with Not_found ->
user_err ~hdr:"lookup_map" (str"map "++qs map++str"not found")
-let protect_red map env sigma c =
- kl (create_clos_infos all env)
- (mk_clos_but (lookup_map map c) (Esubst.subs_id 0) c);;
+let protect_red map env sigma c0 =
+ let evars ev = Evarutil.safe_evar_value sigma ev in
+ let c = EConstr.Unsafe.to_constr c0 in
+ EConstr.of_constr (kl (create_clos_infos ~evars all env)
+ (mk_clos_but (lookup_map map sigma c0) (Esubst.subs_id 0) c));;
let protect_tac map =
Tactics.reduct_option (protect_red map,DEFAULTcast) None
@@ -97,9 +102,10 @@ let protect_tac_in map id =
let closed_term t l =
let open Quote_plugin in
+ Proofview.tclEVARMAP >>= fun sigma ->
let l = List.map Universes.constr_of_global l in
let cs = List.fold_right Quote.ConstrSet.add l Quote.ConstrSet.empty in
- if Quote.closed_under cs t then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (mt())
+ if Quote.closed_under sigma cs t then Proofview.tclUNIT () else Tacticals.New.tclFAIL 0 (mt())
(* TACTIC EXTEND echo
| [ "echo" constr(t) ] ->
@@ -136,14 +142,16 @@ let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term"
let ic c =
let env = Global.env() in
let sigma = Evd.from_env env in
- Constrintern.interp_open_constr env sigma c
+ let sigma, c = Constrintern.interp_open_constr env sigma c in
+ (sigma, c)
let ic_unsafe c = (*FIXME remove *)
let env = Global.env() in
let sigma = Evd.from_env env in
- fst (Constrintern.interp_constr env sigma c)
+ EConstr.of_constr (fst (Constrintern.interp_constr env sigma c))
let decl_constant na ctx c =
+ let open Constr in
let vars = Universes.universes_of_constr c in
let ctx = Universes.restrict_universe_context (Univ.ContextSet.of_context ctx) vars in
mkConst(declare_constant (Id.of_string na)
@@ -172,11 +180,11 @@ let ltac_apply (f : Value.t) (args: Tacinterp.Value.t list) =
let dummy_goal env sigma =
let (gl,_,sigma) =
- Goal.V82.mk_goal sigma (named_context_val env) mkProp Evd.Store.empty in
+ Goal.V82.mk_goal sigma (named_context_val env) EConstr.mkProp Evd.Store.empty in
{Evd.it = gl; Evd.sigma = sigma}
let constr_of v = match Value.to_constr v with
- | Some c -> c
+ | Some c -> EConstr.Unsafe.to_constr c
| None -> failwith "Ring.exec_tactic: anomaly"
let tactic_res = ref [||]
@@ -211,7 +219,8 @@ let exec_tactic env evd n f args =
let gl = dummy_goal env evd in
let gls = Proofview.V82.of_tactic (Tacinterp.eval_tactic_ist ist (ltac_call f (args@[getter]))) gl in
let evd, nf = Evarutil.nf_evars_and_universes (Refiner.project gls) in
- Array.map (fun x -> nf (constr_of x)) !tactic_res, snd (Evd.universe_context evd)
+ let nf c = nf (constr_of c) in
+ Array.map nf !tactic_res, snd (Evd.universe_context evd)
let stdlib_modules =
[["Coq";"Setoids";"Setoid"];
@@ -221,7 +230,7 @@ let stdlib_modules =
]
let coq_constant c =
- lazy (Coqlib.gen_constant_in_modules "Ring" stdlib_modules c)
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "Ring" stdlib_modules c))
let coq_reference c =
lazy (Coqlib.gen_reference_in_modules "Ring" stdlib_modules c)
@@ -239,19 +248,19 @@ let plapp evd f args =
let fc = Evarutil.e_new_global evd (Lazy.force f) in
mkApp(fc,args)
-let dest_rel0 t =
- match kind_of_term t with
+let dest_rel0 sigma t =
+ match EConstr.kind sigma t with
| App(f,args) when Array.length args >= 2 ->
let rel = mkApp(f,Array.sub args 0 (Array.length args - 2)) in
- if closed0 rel then
+ if closed0 sigma rel then
(rel,args.(Array.length args - 2),args.(Array.length args - 1))
else error "ring: cannot find relation (not closed)"
| _ -> error "ring: cannot find relation"
-let rec dest_rel t =
- match kind_of_term t with
- | Prod(_,_,c) -> dest_rel c
- | _ -> dest_rel0 t
+let rec dest_rel sigma t =
+ match EConstr.kind sigma t with
+ | Prod(_,_,c) -> dest_rel sigma c
+ | _ -> dest_rel0 sigma t
(****************************************************************************)
(* Library linking *)
@@ -266,7 +275,7 @@ let plugin_modules =
]
let my_constant c =
- lazy (Coqlib.gen_constant_in_modules "Ring" plugin_modules c)
+ lazy (EConstr.of_constr (Coqlib.gen_constant_in_modules "Ring" plugin_modules c))
let my_reference c =
lazy (Coqlib.gen_reference_in_modules "Ring" plugin_modules c)
@@ -310,13 +319,13 @@ let coq_mkhypo = my_reference "mkhypo"
let coq_hypo = my_reference "hypo"
(* Equality: do not evaluate but make recursive call on both sides *)
-let map_with_eq arg_map c =
- let (req,_,_) = dest_rel c in
+let map_with_eq arg_map sigma c =
+ let (req,_,_) = dest_rel sigma c in
interp_map
- ((global_head_of_constr req,(function -1->Prot|_->Rec))::
+ ((global_head_of_constr sigma req,(function -1->Prot|_->Rec))::
List.map (fun (c,map) -> (Lazy.force c,map)) arg_map)
-let map_without_eq arg_map _ =
+let map_without_eq arg_map _ _ =
interp_map (List.map (fun (c,map) -> (Lazy.force c,map)) arg_map)
let _ = add_map "ring"
@@ -337,6 +346,8 @@ let _ = add_map "ring"
(****************************************************************************)
(* Ring database *)
+let pr_constr c = pr_econstr c
+
module Cmap = Map.Make(Constr)
let from_carrier = Summary.ref Cmap.empty ~name:"ring-tac-carrier-table"
@@ -355,7 +366,7 @@ let find_ring_structure env sigma l =
(str"arguments of ring_simplify do not have all the same type")
in
List.iter check cl';
- (try ring_for_carrier ty
+ (try ring_for_carrier (EConstr.to_constr sigma ty)
with Not_found ->
user_err ~hdr:"ring"
(str"cannot find a declared ring structure over"++
@@ -382,7 +393,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 &&
- eq_constr set' th.ring_setoid &&
+ Term.eq_constr set' th.ring_setoid &&
ext' == th.ring_ext &&
morph' == th.ring_morph &&
th' == th.ring_th &&
@@ -488,8 +499,8 @@ let op_smorph r add mul req m1 m2 =
(* (setoid,op_morph) *)
let ring_equality env evd (r,add,mul,opp,req) =
- match kind_of_term req with
- | App (f, [| _ |]) when eq_constr_nounivs f (Lazy.force coq_eq) ->
+ match EConstr.kind !evd req with
+ | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) ->
let setoid = plapp evd coq_eq_setoid [|r|] in
let op_morph =
match opp with
@@ -543,15 +554,15 @@ let build_setoid_params env evd r add mul opp req eqth =
let dest_ring env sigma th_spec =
let th_typ = Retyping.get_type_of env sigma th_spec in
- match kind_of_term th_typ with
+ match EConstr.kind sigma th_typ with
App(f,[|r;zero;one;add;mul;sub;opp;req|])
- when eq_constr_nounivs f (Lazy.force coq_almost_ring_theory) ->
+ when eq_constr_nounivs sigma f (Lazy.force coq_almost_ring_theory) ->
(None,r,zero,one,add,mul,Some sub,Some opp,req)
| App(f,[|r;zero;one;add;mul;req|])
- when eq_constr_nounivs f (Lazy.force coq_semi_ring_theory) ->
+ when eq_constr_nounivs sigma f (Lazy.force coq_semi_ring_theory) ->
(Some true,r,zero,one,add,mul,None,None,req)
| App(f,[|r;zero;one;add;mul;sub;opp;req|])
- when eq_constr_nounivs f (Lazy.force coq_ring_theory) ->
+ when eq_constr_nounivs sigma f (Lazy.force coq_ring_theory) ->
(Some false,r,zero,one,add,mul,Some sub,Some opp,req)
| _ -> error "bad ring structure"
@@ -584,6 +595,7 @@ let make_hyp_list env evd lH =
(plapp evd coq_nil [|carrier|])
in
let l' = Typing.e_solve_evars env evd l in
+ let l' = EConstr.Unsafe.to_constr l' in
Evarutil.nf_evars_universes !evd l'
let interp_power env evd pow =
@@ -619,7 +631,7 @@ let interp_div env evd div =
plapp evd coq_Some [|carrier;spec|]
(* Same remark on ill-typed terms ... *)
-let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div =
+let add_theory0 name (sigma, rth) eqth morphth cst_tac (pre,post) power sign div =
check_required_library (cdir@["Ring_base"]);
let env = Global.env() in
let (kind,r,zero,one,add,mul,sub,opp,req) = dest_ring env sigma rth in
@@ -649,6 +661,9 @@ let add_theory name (sigma,rth) eqth morphth cst_tac (pre,post) power sign div =
match post with
Some t -> Tacintern.glob_tactic t
| _ -> TacId [] in
+ let r = EConstr.to_constr sigma r in
+ let req = EConstr.to_constr sigma req in
+ let sth = EConstr.to_constr sigma sth in
let _ =
Lib.add_leaf name
(theory_to_obj
@@ -696,13 +711,18 @@ let process_ring_mods l =
let k = match !kind with Some k -> k | None -> Abstract in
(k, !set, !cst_tac, !pre, !post, !power, !sign, !div)
+let add_theory id rth l =
+ let (sigma, rth) = ic rth in
+ let (k,set,cst,pre,post,power,sign, div) = process_ring_mods l in
+ add_theory0 id (sigma, rth) set k cst (pre,post) power sign div
+
(*****************************************************************************)
(* The tactics consist then only in a lookup in the ring database and
call the appropriate ltac. *)
-let make_args_list rl t =
+let make_args_list sigma rl t =
match rl with
- | [] -> let (_,t1,t2) = dest_rel0 t in [t1;t2]
+ | [] -> let (_,t1,t2) = dest_rel0 sigma t in [t1;t2]
| _ -> rl
let make_term_list env evd carrier rl =
@@ -711,7 +731,7 @@ let make_term_list env evd carrier rl =
(plapp evd coq_nil [|carrier|])
in Typing.e_solve_evars env evd l
-let carg = Tacinterp.Value.of_constr
+let carg c = Tacinterp.Value.of_constr (EConstr.of_constr c)
let tacarg expr =
Tacinterp.Value.of_closure (Tacinterp.default_ist ()) expr
@@ -735,10 +755,10 @@ let ring_lookup (f : Value.t) lH rl t =
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
try (* find_ring_strucure can raise an exception *)
+ let rl = make_args_list sigma rl t in
let evdref = ref sigma in
- let rl = make_args_list rl t in
let e = find_ring_structure env sigma rl in
- let rl = carg (make_term_list env evdref e.ring_carrier rl) in
+ let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in
let lH = carg (make_hyp_list env evdref lH) in
let ring = ltac_ring_structure e in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (ring@[lH;rl]))
@@ -801,21 +821,22 @@ let af_ar = my_reference"AF_AR"
let f_r = my_reference"F_R"
let sf_sr = my_reference"SF_SR"
let dest_field env evd th_spec =
+ let open Termops in
let th_typ = Retyping.get_type_of env !evd th_spec in
- match kind_of_term th_typ with
+ match EConstr.kind !evd th_typ with
| App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|])
- when is_global (Lazy.force afield_theory) f ->
+ when is_global !evd (Lazy.force afield_theory) f ->
let rth = plapp evd af_ar
[|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in
(None,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth)
| App(f,[|r;zero;one;add;mul;sub;opp;div;inv;req|])
- when is_global (Lazy.force field_theory) f ->
+ when is_global !evd (Lazy.force field_theory) f ->
let rth =
plapp evd f_r
[|r;zero;one;add;mul;sub;opp;div;inv;req;th_spec|] in
(Some false,r,zero,one,add,mul,Some sub,Some opp,div,inv,req,rth)
| App(f,[|r;zero;one;add;mul;div;inv;req|])
- when is_global (Lazy.force sfield_theory) f ->
+ when is_global !evd (Lazy.force sfield_theory) f ->
let rth = plapp evd sf_sr
[|r;zero;one;add;mul;div;inv;req;th_spec|] in
(Some true,r,zero,one,add,mul,None,None,div,inv,req,rth)
@@ -838,7 +859,7 @@ let find_field_structure env sigma l =
(str"arguments of field_simplify do not have all the same type")
in
List.iter check cl';
- (try field_for_carrier ty
+ (try field_for_carrier (EConstr.to_constr sigma ty)
with Not_found ->
user_err ~hdr:"field"
(str"cannot find a declared field structure over"++
@@ -895,9 +916,11 @@ let ftheory_to_obj : field_info -> obj =
classify_function = (fun x -> Substitute x) }
let field_equality evd r inv req =
- match kind_of_term req with
- | App (f, [| _ |]) when eq_constr_nounivs f (Lazy.force coq_eq) ->
- mkApp(Universes.constr_of_global (Coqlib.build_coq_eq_data()).congr,[|r;r;inv|])
+ match EConstr.kind !evd req with
+ | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) ->
+ let c = Universes.constr_of_global (Coqlib.build_coq_eq_data()).congr in
+ let c = EConstr.of_constr c in
+ mkApp(c,[|r;r;inv|])
| _ ->
let _setoid = setoid_of_relation (Global.env ()) evd r req in
let signature = [Some (r,Some req)],Some(r,Some req) in
@@ -907,15 +930,17 @@ let field_equality evd r inv req =
error "field inverse should be declared as a morphism" in
inv_m_lem
-let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power sign odiv =
+let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv =
+ let open Constr in
check_required_library (cdir@["Field_tac"]);
+ let (sigma,fth) = ic fth in
let env = Global.env() in
let evd = ref sigma in
let (kind,r,zero,one,add,mul,sub,opp,div,inv,req,rth) =
dest_field env evd fth in
let (sth,ext) = build_setoid_params env evd r add mul opp req eqth in
let eqth = Some(sth,ext) in
- let _ = add_theory name (!evd,rth) eqth morphth cst_tac (None,None) power sign odiv in
+ let _ = add_theory0 name (!evd,rth) eqth morphth cst_tac (None,None) power sign odiv in
let (pow_tac, pspec) = interp_power env evd power in
let sspec = interp_sign env evd sign in
let dspec = interp_div env evd odiv in
@@ -930,7 +955,7 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power
let lemma4 = params.(6) in
let cond_lemma =
match inj with
- | Some thm -> mkApp(params.(8),[|thm|])
+ | Some thm -> mkApp(params.(8),[|EConstr.to_constr sigma thm|])
| None -> params.(7) in
let lemma1 = decl_constant (Id.to_string name^"_field_lemma1")
ctx lemma1 in
@@ -952,6 +977,8 @@ let add_field_theory name (sigma,fth) eqth morphth cst_tac inj (pre,post) power
match post with
Some t -> Tacintern.glob_tactic t
| _ -> TacId [] in
+ let r = EConstr.to_constr sigma r in
+ let req = EConstr.to_constr sigma req in
let _ =
Lib.add_leaf name
(ftheory_to_obj
@@ -991,6 +1018,10 @@ let process_field_mods l =
let k = match !kind with Some k -> k | None -> Abstract in
(k, !set, !inj, !cst_tac, !pre, !post, !power, !sign, !div)
+let add_field_theory id t mods =
+ let (k,set,inj,cst_tac,pre,post,power,sign,div) = process_field_mods mods in
+ add_field_theory0 id t set k cst_tac inj (pre,post) power sign div
+
let ltac_field_structure e =
let req = carg e.field_req in
let cst_tac = tacarg e.field_cst_tac in
@@ -1010,10 +1041,10 @@ let field_lookup (f : Value.t) lH rl t =
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
try
+ let rl = make_args_list sigma rl t in
let evdref = ref sigma in
- let rl = make_args_list rl t in
let e = find_field_structure env sigma rl in
- let rl = carg (make_term_list env evdref e.field_carrier rl) in
+ let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in
let lH = carg (make_hyp_list env evdref lH) in
let field = ltac_field_structure e in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (ltac_apply f (field@[lH;rl]))
diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli
index f417c87cd..4367d021c 100644
--- a/plugins/setoid_ring/newring.mli
+++ b/plugins/setoid_ring/newring.mli
@@ -8,6 +8,7 @@
open Names
open Constr
+open EConstr
open Libnames
open Globnames
open Constrexpr
@@ -19,28 +20,12 @@ val protect_tac_in : string -> Id.t -> unit Proofview.tactic
val protect_tac : string -> unit Proofview.tactic
-val closed_term : constr -> global_reference list -> unit Proofview.tactic
-
-val process_ring_mods :
- constr_expr ring_mod list ->
- constr coeff_spec * (constr * constr) option *
- cst_tac_spec option * raw_tactic_expr option *
- raw_tactic_expr option *
- (cst_tac_spec * constr_expr) option *
- constr_expr option * constr_expr option
+val closed_term : EConstr.constr -> global_reference list -> unit Proofview.tactic
val add_theory :
Id.t ->
- Evd.evar_map * constr ->
- (constr * constr) option ->
- constr coeff_spec ->
- cst_tac_spec option ->
- raw_tactic_expr option * raw_tactic_expr option ->
- (cst_tac_spec * constr_expr) option ->
- constr_expr option ->
- constr_expr option -> unit
-
-val ic : constr_expr -> Evd.evar_map * constr
+ constr_expr ->
+ constr_expr ring_mod list -> unit
val from_name : ring_info Spmap.t ref
@@ -49,26 +34,10 @@ val ring_lookup :
constr list ->
constr list -> constr -> unit Proofview.tactic
-val process_field_mods :
- constr_expr field_mod list ->
- constr coeff_spec *
- (constr * constr) option * constr option *
- cst_tac_spec option * raw_tactic_expr option *
- raw_tactic_expr option *
- (cst_tac_spec * constr_expr) option *
- constr_expr option * constr_expr option
-
val add_field_theory :
Id.t ->
- Evd.evar_map * constr ->
- (constr * constr) option ->
- constr coeff_spec ->
- cst_tac_spec option ->
- constr option ->
- raw_tactic_expr option * raw_tactic_expr option ->
- (cst_tac_spec * constr_expr) option ->
- constr_expr option ->
- constr_expr option -> unit
+ constr_expr ->
+ constr_expr field_mod list -> unit
val field_from_name : field_info Spmap.t ref
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 3d0232d94..f3555ebc4 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -179,6 +179,9 @@ let mk_lterm = mk_term ' '
let pf_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty
+let nf_evar sigma c =
+ EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr c))
+
(* }}} *)
(** Profiling {{{ *************************************************************)
@@ -306,7 +309,7 @@ let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise
let unif_HO_args env ise0 pa i ca =
let n = Array.length pa in
let rec loop ise j =
- if j = n then ise else loop (unif_HO env ise pa.(j) ca.(i + j)) (j + 1) in
+ if j = n then ise else loop (unif_HO env ise (EConstr.of_constr pa.(j)) (EConstr.of_constr ca.(i + j))) (j + 1) in
loop ise0 0
(* FO unification should boil down to calling w_unify with no_delta, but *)
@@ -333,10 +336,11 @@ let flags_FO =
(Unification.default_no_delta_unify_flags ()).Unification.resolve_evars
}
let unif_FO env ise p c =
- Unification.w_unify env ise Reduction.CONV ~flags:flags_FO p c
+ Unification.w_unify env ise Reduction.CONV ~flags:flags_FO (EConstr.of_constr p) (EConstr.of_constr c)
(* Perform evar substitution in main term and prune substitution. *)
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
| Evar ex ->
@@ -353,7 +357,7 @@ let nf_open_term sigma0 ise c =
| Evar_defined c' -> s' := Evd.define k (nf c') !s'
| _ -> () in
let c' = nf c in let _ = Evd.fold copy_def sigma0 () in
- !s', Evd.evar_universe_context s, c'
+ !s', Evd.evar_universe_context s, EConstr.of_constr c'
let unif_end env sigma0 ise0 pt ok =
let ise = Evarconv.solve_unif_constraints_with_heuristics env ise0 in
@@ -428,7 +432,7 @@ type tpattern = {
up_a : constr array;
up_t : constr; (* equation proof term or matched term *)
up_dir : ssrdir; (* direction of the rule *)
- up_ok : constr -> evar_map -> bool; (* progess test for rewrite *)
+ up_ok : constr -> evar_map -> bool; (* progress test for rewrite *)
}
let all_ok _ _ = true
@@ -483,7 +487,9 @@ let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 =
(* p_origin can be passed to obtain a better error message *)
let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p =
let k, f, a =
- let f, a = Reductionops.whd_betaiota_stack ise p in
+ 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
| Const (p,_) ->
let np = proj_nparams p in
@@ -640,13 +646,14 @@ let match_upats_FO upats env sigma0 ise orig_c =
| _ -> unif_FO env ise u.up_FO c' in
let ise' = (* Unify again using HO to assign evars *)
let p = mkApp (u.up_f, u.up_a) in
- try unif_HO env ise p c' with _ -> raise NoMatch in
+ try unif_HO env ise (EConstr.of_constr p) (EConstr.of_constr c') with e when CErrors.noncritical e -> raise NoMatch in
let lhs = mkSubApp f i a in
- let pt' = unif_end env sigma0 ise' u.up_t (u.up_ok lhs) in
+ let pt' = unif_end env sigma0 ise' (EConstr.of_constr u.up_t) (u.up_ok lhs) in
+ let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in
raise (FoundUnif (ungen_upat lhs pt' u))
with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u
| Not_found -> CErrors.anomaly (str"incomplete ise in match_upats_FO")
- | _ -> () in
+ | e when CErrors.noncritical e -> () in
List.iter one_match fpats
done;
iter_constr_LR loop f; Array.iter loop a in
@@ -659,7 +666,7 @@ let match_upats_FO upats env sigma0 ise c =
let match_upats_HO ~on_instance upats env sigma0 ise c =
- let dont_impact_evars = dont_impact_evars_in c in
+ let dont_impact_evars = dont_impact_evars_in c in
let it_did_match = ref false in
let failed_because_of_TC = ref false in
let rec aux upats env sigma0 ise c =
@@ -681,16 +688,17 @@ let match_upats_HO ~on_instance upats env sigma0 ise c =
| KpatLet ->
let x, v, t, b = destLetIn f in
let _, pv, _, pb = destLetIn u.up_f in
- let ise' = unif_HO env ise pv v in
+ let ise' = unif_HO env ise (EConstr.of_constr pv) (EConstr.of_constr v) in
unif_HO
(Environ.push_rel (Context.Rel.Declaration.LocalAssum(x, t)) env)
- ise' pb b
+ ise' (EConstr.of_constr pb) (EConstr.of_constr b)
| KpatFlex | KpatProj _ ->
- unif_HO env ise u.up_f (mkSubApp f (i - Array.length u.up_a) a)
- | _ -> unif_HO env ise u.up_f f in
+ unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr(mkSubApp f (i - Array.length u.up_a) a))
+ | _ -> unif_HO env ise (EConstr.of_constr u.up_f) (EConstr.of_constr f) in
let ise'' = unif_HO_args env ise' u.up_a (i - Array.length u.up_a) a in
let lhs = mkSubApp f i a in
- let pt' = unif_end env sigma0 ise'' u.up_t (u.up_ok lhs) in
+ let pt' = unif_end env sigma0 ise'' (EConstr.of_constr u.up_t) (u.up_ok lhs) in
+ let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in
on_instance (ungen_upat lhs pt' u)
with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u
| NoProgress -> it_did_match := true
@@ -715,7 +723,7 @@ let match_upats_HO ~on_instance upats env sigma0 ise c =
let fixed_upat = function
| {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false
-| {up_t = t} -> not (occur_existential t)
+| {up_t = t} -> not (occur_existential Evd.empty (EConstr.of_constr t)) (** FIXME *)
let do_once r f = match !r with Some _ -> () | None -> r := Some (f ())
@@ -729,13 +737,13 @@ let assert_done_multires r =
r := Some (n+1,xs);
try List.nth xs n with Failure _ -> raise NoMatch
-type subst = Environ.env -> Term.constr -> Term.constr -> int -> Term.constr
+type subst = Environ.env -> constr -> constr -> int -> constr
type find_P =
- Environ.env -> Term.constr -> int ->
+ Environ.env -> constr -> int ->
k:subst ->
- Term.constr
+ constr
type conclude = unit ->
- Term.constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * Term.constr)
+ constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * constr)
(* upats_origin makes a better error message only *)
let mk_tpattern_matcher ?(all_instances=false)
@@ -791,13 +799,13 @@ let on_instance, instances =
let rec uniquize = function
| [] -> []
| (sigma,_,{ up_f = f; up_a = a; up_t = t } as x) :: xs ->
- let t = Reductionops.nf_evar sigma t in
- let f = Reductionops.nf_evar sigma f in
- let a = Array.map (Reductionops.nf_evar sigma) a in
+ let t = nf_evar sigma t in
+ let f = nf_evar sigma f in
+ let a = Array.map (nf_evar sigma) a in
let neq (sigma1,_,{ up_f = f1; up_a = a1; up_t = t1 }) =
- let t1 = Reductionops.nf_evar sigma1 t1 in
- let f1 = Reductionops.nf_evar sigma1 f1 in
- let a1 = Array.map (Reductionops.nf_evar sigma1) a1 in
+ 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
x :: uniquize (List.filter neq xs) in
@@ -846,8 +854,11 @@ let rec uniquize = function
| Context.Rel.Declaration.LocalAssum _ as x -> x
| Context.Rel.Declaration.LocalDef (x,_,y) ->
Context.Rel.Declaration.LocalAssum(x,y) in
- Environ.push_rel ctx_item env, h' + 1 in
- let f' = map_constr_with_binders_left_to_right inc_h subst_loop acc f in
+ EConstr.push_rel ctx_item env, h' + 1 in
+ let self acc c = EConstr.of_constr (subst_loop acc (EConstr.Unsafe.to_constr c)) in
+ let f = EConstr.of_constr f in
+ let f' = map_constr_with_binders_left_to_right sigma inc_h self acc f in
+ let f' = EConstr.Unsafe.to_constr f' in
mkApp (f', Array.map_left (subst_loop acc) a) in
subst_loop (env,h) c) : find_P),
((fun () ->
@@ -902,7 +913,7 @@ let pr_pattern_aux pr_constr = function
| E_As_X_In_T (e,x,t) ->
pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t
let pp_pattern (sigma, p) =
- pr_pattern_aux (fun t -> pr_constr_pat (pi3 (nf_open_term sigma sigma t))) p
+ pr_pattern_aux (fun t -> pr_constr_pat (EConstr.Unsafe.to_constr (pi3 (nf_open_term sigma sigma (EConstr.of_constr t))))) p
let pr_cpattern = pr_term
let pr_rpattern _ _ _ = pr_pattern
@@ -1006,7 +1017,7 @@ type occ = (bool * int list) option
type rpattern = (cpattern, cpattern) ssrpattern
let pr_rpattern = pr_pattern
-type pattern = Evd.evar_map * (Term.constr, Term.constr) ssrpattern
+type pattern = Evd.evar_map * (constr, constr) ssrpattern
let id_of_cpattern = function
@@ -1038,7 +1049,7 @@ let interp_constr = interp_wit wit_constr
let interp_open_constr ist gl gc =
interp_wit wit_open_constr ist gl gc
let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c
-let interp_term ist gl (_, c) = (interp_open_constr ist gl c)
+let interp_term ist gl (_, c) = on_snd EConstr.Unsafe.to_constr (interp_open_constr ist gl c)
let pr_ssrterm _ _ _ = pr_term
let input_ssrtermkind strm = match stream_nth 0 strm with
| Tok.KEYWORD "(" -> '('
@@ -1144,7 +1155,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
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
- aux [] (Evarutil.nf_evar sigma rp) in
+ aux [] (nf_evar sigma rp) in
let sigma =
List.fold_left (fun sigma e ->
if Evd.is_defined sigma e then sigma else (* clear may be recursive *)
@@ -1201,7 +1212,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
let sigma, rp = interp_term ist gl rp in
let _, h, _, rp = destLetIn rp in
let sigma = cleanup_XinE h x rp sigma in
- let rp = subst1 h (Evarutil.nf_evar sigma rp) in
+ let rp = subst1 h (nf_evar sigma rp) in
sigma, mk h rp
| E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) ->
let mk e x p =
@@ -1210,7 +1221,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
let sigma, rp = interp_term ist gl rp in
let _, h, _, rp = destLetIn rp in
let sigma = cleanup_XinE h x rp sigma in
- let rp = subst1 h (Evarutil.nf_evar sigma rp) in
+ let rp = subst1 h (nf_evar sigma rp) in
let sigma, e = interp_term ist (re_sig (sig_it gl) sigma) e in
sigma, mk e h rp
;;
@@ -1226,7 +1237,7 @@ let noindex = Some(false,[])
(* calls do_subst on every sub-term identified by (pattern,occ) *)
let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
- let fs sigma x = Reductionops.nf_evar sigma x in
+ let fs sigma x = nf_evar sigma x in
let pop_evar sigma e p =
let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in
let e_body = match e_body with Evar_defined c -> c
@@ -1263,7 +1274,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in
let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in
let concl = find_T env0 concl0 1 (fun env c _ h ->
- let p_sigma = unify_HO env (create_evar_defs sigma) c p in
+ let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in
let sigma, e_body = pop_evar p_sigma ex p in
fs p_sigma (find_X env (fs sigma p) h
(fun env _ -> do_subst env e_body))) in
@@ -1279,7 +1290,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in
let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in
let concl = find_T env0 concl0 1 (fun env c _ h ->
- let p_sigma = unify_HO env (create_evar_defs sigma) c p in
+ let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in
let sigma, e_body = pop_evar p_sigma ex p in
fs p_sigma (find_X env (fs sigma p) h (fun env c _ h ->
find_E env e_body h do_subst))) in
@@ -1289,17 +1300,17 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let p, e = fs sigma p, fs sigma e in
let ex = ex_value hole in
let rp =
- let e_sigma = unify_HO env0 sigma hole e in
+ let e_sigma = unify_HO env0 sigma (EConstr.of_constr hole) (EConstr.of_constr e) in
e_sigma, fs e_sigma p in
let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in
let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in
let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in
let find_X, end_X = mk_tpattern_matcher sigma occ holep in
let concl = find_TE env0 concl0 1 (fun env c _ h ->
- let p_sigma = unify_HO env (create_evar_defs sigma) c p in
+ let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in
let sigma, e_body = pop_evar p_sigma ex p in
fs p_sigma (find_X env (fs sigma p) h (fun env c _ h ->
- let e_sigma = unify_HO env sigma e_body e in
+ let e_sigma = unify_HO env sigma (EConstr.of_constr e_body) (EConstr.of_constr e) in
let e_body = fs e_sigma e in
do_subst env e_body e_body h))) in
let _ = end_X () in let _ = end_TE () in
@@ -1313,7 +1324,7 @@ let redex_of_pattern ?(resolve_typeclasses=false) env (sigma, p) =
let sigma =
if not resolve_typeclasses then sigma
else Typeclasses.resolve_typeclasses ~fail:false env sigma in
- Reductionops.nf_evar sigma e, Evd.evar_universe_context sigma
+ nf_evar sigma e, Evd.evar_universe_context sigma
let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h =
let do_make_rel, occ =
@@ -1335,13 +1346,15 @@ let mk_tpattern ?p_origin env sigma0 sigma_t f dir c =
;;
let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h =
+ let p = EConstr.Unsafe.to_constr p in
+ let concl = EConstr.Unsafe.to_constr concl in
let ise = create_evar_defs sigma in
- let ise, u = mk_tpattern env sigma0 (ise,t) ok L2R p in
+ let ise, u = mk_tpattern env sigma0 (ise,EConstr.Unsafe.to_constr t) ok L2R p in
let find_U, end_U =
mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in
let concl = find_U env concl h (fun _ _ _ -> mkRel) in
let rdx, _, (sigma, uc, p) = end_U () in
- sigma, uc, p, concl, rdx
+ sigma, uc, EConstr.of_constr p, EConstr.of_constr concl, EConstr.of_constr rdx
let fill_occ_term env cl occ sigma0 (sigma, t) =
try
@@ -1354,7 +1367,7 @@ let fill_occ_term env cl occ sigma0 (sigma, t) =
if sigma' != sigma0 then raise NoMatch
else cl, (Evd.merge_universe_context sigma' uc, t')
with _ ->
- errorstrm (str "partial term " ++ pr_constr_pat t
+ errorstrm (str "partial term " ++ pr_constr_pat (EConstr.Unsafe.to_constr t)
++ str " does not match any subterm of the goal")
let pf_fill_occ_term gl occ t =
@@ -1394,10 +1407,13 @@ let ssrpatterntac _ist (arg_ist,arg) gl =
let pat = interp_rpattern arg_ist gl arg in
let sigma0 = project gl in
let concl0 = pf_concl gl in
+ let concl0 = EConstr.Unsafe.to_constr concl0 in
let (t, uc), concl_x =
fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in
+ let t = EConstr.of_constr t in
+ let concl_x = EConstr.of_constr concl_x in
let gl, tty = pf_type_of gl t in
- let concl = mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in
+ let concl = EConstr.mkLetIn (Name (id_of_string "selected"), t, tty, concl_x) in
Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl
(* Register "ssrpattern" tactic *)
@@ -1420,6 +1436,7 @@ let ssrinstancesof ist arg gl =
let ok rhs lhs ise = true in
(* not (Term.eq_constr 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
let pat = match cpat with T x -> x | _ -> errorstrm (str"Not supported") in
let etpat, tpat = mk_tpattern env sigma (sigma0,pat) (ok pat) L2R pat in
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
index 288a04e60..894cdb943 100644
--- a/plugins/ssrmatching/ssrmatching.mli
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -194,7 +194,7 @@ val mk_tpattern_matcher :
(* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns
* the conclusion of [gl] where [occ] occurrences of [t] have been replaced
* by [Rel 1] and the instance of [t] *)
-val pf_fill_occ_term : goal sigma -> occ -> evar_map * constr -> constr * constr
+val pf_fill_occ_term : goal sigma -> occ -> evar_map * EConstr.t -> EConstr.t * EConstr.t
(* It may be handy to inject a simple term into the first form of cpattern *)
val cpattern_of_term : char * glob_constr_and_expr -> cpattern
@@ -216,8 +216,8 @@ val assert_done : 'a option ref -> 'a
[solve_unif_constraints_with_heuristics] and [resolve_typeclasses].
In case of failure they raise [NoMatch] *)
-val unify_HO : env -> evar_map -> constr -> constr -> evar_map
-val pf_unify_HO : goal sigma -> constr -> constr -> goal sigma
+val unify_HO : env -> evar_map -> EConstr.constr -> EConstr.constr -> evar_map
+val pf_unify_HO : goal sigma -> EConstr.constr -> EConstr.constr -> goal sigma
(** Some more low level functions needed to implement the full SSR language
on top of the former APIs *)
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 63c2dde18..c5cf74ccf 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -6,18 +6,21 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Declarations
open Inductiveops
-open Environ
open Reductionops
open Type_errors
open Glob_term
@@ -99,11 +102,12 @@ let make_anonymous_patvars n =
let relocate_rel n1 n2 k j = if Int.equal j (n1 + k) then n2+k else j
-let rec relocate_index n1 n2 k t = match kind_of_term t with
+let rec relocate_index sigma n1 n2 k t =
+ match EConstr.kind sigma t with
| Rel j when Int.equal j (n1 + k) -> mkRel (n2+k)
| Rel j when j < n1+k -> t
| Rel j when j > n1+k -> t
- | _ -> map_constr_with_binders succ (relocate_index n1 n2) k t
+ | _ -> EConstr.map_with_binders sigma succ (relocate_index sigma n1 n2) k t
(**********************************************************************)
(* Structures used in compiling pattern-matching *)
@@ -138,7 +142,7 @@ type tomatch_status =
| Pushed of (bool*((constr * tomatch_type) * int list * Name.t))
| Alias of (bool*(Name.t * constr * (constr * types)))
| NonDepAlias
- | Abstract of int * Context.Rel.Declaration.t
+ | Abstract of int * rel_declaration
type tomatch_stack = tomatch_status list
@@ -275,6 +279,7 @@ let rec find_row_ind = function
let inductive_template evdref env tmloc ind =
let indu = evd_comb1 (Evd.fresh_inductive_instance env) evdref ind in
let arsign = inductive_alldecls_env env indu in
+ let indu = on_snd EInstance.make indu in
let hole_source i = match tmloc with
| Some loc -> (loc, Evar_kinds.TomatchTypeParameter (ind,i))
| None -> (Loc.ghost, Evar_kinds.TomatchTypeParameter (ind,i)) in
@@ -283,10 +288,12 @@ let inductive_template evdref env tmloc ind =
(fun decl (subst,evarl,n) ->
match decl with
| LocalAssum (na,ty) ->
+ let ty = EConstr.of_constr ty in
let ty' = substl subst ty in
let e = e_new_evar env evdref ~src:(hole_source n) ty' in
(e::subst,e::evarl,n+1)
| LocalDef (na,b,ty) ->
+ let b = EConstr.of_constr b in
(substl subst b::subst,evarl,n+1))
arsign ([],[],1) in
applist (mkIndU indu,List.rev evarl)
@@ -310,9 +317,9 @@ let inh_coerce_to_ind evdref env loc ty tyi =
un inductif cela doit être égal *)
if not (e_cumul env evdref expected_typ ty) then evdref := sigma
-let binding_vars_of_inductive = function
+let binding_vars_of_inductive sigma = function
| NotInd _ -> []
- | IsInd (_,IndType(_,realargs),_) -> List.filter isRel realargs
+ | IsInd (_,IndType(_,realargs),_) -> List.filter (isRel sigma) realargs
let extract_inductive_data env sigma decl =
match decl with
@@ -320,7 +327,7 @@ let extract_inductive_data env sigma decl =
let tmtyp =
try try_find_ind env sigma t None
with Not_found -> NotInd (None,t) in
- let tmtypvars = binding_vars_of_inductive tmtyp in
+ let tmtypvars = binding_vars_of_inductive sigma tmtyp in
(tmtyp,tmtypvars)
| LocalDef (_,_,t) ->
(NotInd (None, t), [])
@@ -390,7 +397,7 @@ let adjust_tomatch_to_pattern pb ((current,typ),deps,dep) =
| Some (_,(ind,_)) ->
let indt = inductive_template pb.evdref pb.env None ind in
let current =
- if List.is_empty deps && isEvar typ then
+ if List.is_empty deps && isEvar !(pb.evdref) typ then
(* Don't insert coercions if dependent; only solve evars *)
let _ = e_cumul pb.env pb.evdref indt typ in
current
@@ -409,7 +416,7 @@ let map_tomatch_type f = function
| IsInd (t,ind,names) -> IsInd (f t,map_inductive_type f ind,names)
| NotInd (c,t) -> NotInd (Option.map f c, f t)
-let liftn_tomatch_type n depth = map_tomatch_type (liftn n depth)
+let liftn_tomatch_type n depth = map_tomatch_type (Vars.liftn n depth)
let lift_tomatch_type n = liftn_tomatch_type n 1
(**********************************************************************)
@@ -535,19 +542,19 @@ let dependencies_in_pure_rhs nargs eqns =
let deps_columns = matrix_transpose deps_rows in
List.map (List.exists (fun x -> x)) deps_columns
-let dependent_decl a =
+let dependent_decl sigma a =
function
- | LocalAssum (na,t) -> dependent a t
- | LocalDef (na,c,t) -> dependent a t || dependent a c
+ | LocalAssum (na,t) -> dependent sigma a t
+ | LocalDef (na,c,t) -> dependent sigma a t || dependent sigma a c
-let rec dep_in_tomatch n = function
- | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch n l
- | Abstract (_,d) :: l -> dependent_decl (mkRel n) d || dep_in_tomatch (n+1) l
+let rec dep_in_tomatch sigma n = function
+ | (Pushed _ | Alias _ | NonDepAlias) :: l -> dep_in_tomatch sigma n l
+ | Abstract (_,d) :: l -> dependent_decl sigma (mkRel n) d || dep_in_tomatch sigma (n+1) l
| [] -> false
-let dependencies_in_rhs nargs current tms eqns =
- match kind_of_term current with
- | Rel n when dep_in_tomatch n tms -> List.make nargs true
+let dependencies_in_rhs sigma nargs current tms eqns =
+ match EConstr.kind sigma current with
+ | Rel n when dep_in_tomatch sigma n tms -> List.make nargs true
| _ -> dependencies_in_pure_rhs nargs eqns
(* Computing the matrix of dependencies *)
@@ -562,24 +569,24 @@ let dependencies_in_rhs nargs current tms eqns =
[n-2;1], [1] points to [dn] and [n-2] to [d3]
*)
-let rec find_dependency_list tmblock = function
+let rec find_dependency_list sigma tmblock = function
| [] -> []
| (used,tdeps,d)::rest ->
- let deps = find_dependency_list tmblock rest in
- if used && List.exists (fun x -> dependent_decl x d) tmblock
+ let deps = find_dependency_list sigma tmblock rest in
+ if used && List.exists (fun x -> dependent_decl sigma x d) tmblock
then
List.add_set Int.equal
(List.length rest + 1) (List.union Int.equal deps tdeps)
else deps
-let find_dependencies is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist =
- let deps = find_dependency_list (tm::tmtypleaves) nextlist in
+let find_dependencies sigma is_dep_or_cstr_in_rhs (tm,(_,tmtypleaves),d) nextlist =
+ let deps = find_dependency_list sigma (tm::tmtypleaves) nextlist in
if is_dep_or_cstr_in_rhs || not (List.is_empty deps)
then ((true ,deps,d)::nextlist)
else ((false,[] ,d)::nextlist)
-let find_dependencies_signature deps_in_rhs typs =
- let l = List.fold_right2 find_dependencies deps_in_rhs typs [] in
+let find_dependencies_signature sigma deps_in_rhs typs =
+ let l = List.fold_right2 (find_dependencies sigma) deps_in_rhs typs [] in
List.map (fun (_,deps,_) -> deps) l
(* Assume we had terms t1..tq to match in a context xp:Tp,...,x1:T1 |-
@@ -593,31 +600,31 @@ let find_dependencies_signature deps_in_rhs typs =
[relocate_index_tomatch 1 n tomatch] will go the way back.
*)
-let relocate_index_tomatch n1 n2 =
+let relocate_index_tomatch sigma n1 n2 =
let rec genrec depth = function
| [] ->
[]
| Pushed (b,((c,tm),l,na)) :: rest ->
- let c = relocate_index n1 n2 depth c in
- let tm = map_tomatch_type (relocate_index n1 n2 depth) tm in
+ let c = relocate_index sigma n1 n2 depth c in
+ let tm = map_tomatch_type (relocate_index sigma n1 n2 depth) tm in
let l = List.map (relocate_rel n1 n2 depth) l in
Pushed (b,((c,tm),l,na)) :: genrec depth rest
| Alias (initial,(na,c,d)) :: rest ->
(* [c] is out of relocation scope *)
- Alias (initial,(na,c,map_pair (relocate_index n1 n2 depth) d)) :: genrec depth rest
+ Alias (initial,(na,c,map_pair (relocate_index sigma n1 n2 depth) d)) :: genrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: genrec depth rest
| Abstract (i,d) :: rest ->
let i = relocate_rel n1 n2 depth i in
- Abstract (i, RelDecl.map_constr (relocate_index n1 n2 depth) d)
+ Abstract (i, RelDecl.map_constr (fun c -> relocate_index sigma n1 n2 depth c) d)
:: genrec (depth+1) rest in
genrec 0
(* [replace_tomatch n c tomatch] replaces [Rel n] by [c] in [tomatch] *)
-let rec replace_term n c k t =
- if isRel t && Int.equal (destRel t) (n + k) then lift k c
- else map_constr_with_binders succ (replace_term n c) k t
+let rec replace_term sigma n c k t =
+ if isRel sigma t && Int.equal (destRel sigma t) (n + k) then Vars.lift k c
+ else EConstr.map_with_binders sigma succ (replace_term sigma n c) k t
let length_of_tomatch_type_sign na t =
let l = match na with
@@ -628,21 +635,21 @@ let length_of_tomatch_type_sign na t =
| NotInd _ -> l
| IsInd (_, _, names) -> List.length names + l
-let replace_tomatch n c =
+let replace_tomatch sigma n c =
let rec replrec depth = function
| [] -> []
| Pushed (initial,((b,tm),l,na)) :: rest ->
- let b = replace_term n c depth b in
- let tm = map_tomatch_type (replace_term n c depth) tm in
+ let b = replace_term sigma n c depth b in
+ let tm = map_tomatch_type (replace_term sigma n c depth) tm in
List.iter (fun i -> if Int.equal i (n + depth) then anomaly (Pp.str "replace_tomatch")) l;
Pushed (initial,((b,tm),l,na)) :: replrec depth rest
| Alias (initial,(na,b,d)) :: rest ->
(* [b] is out of replacement scope *)
- Alias (initial,(na,b,map_pair (replace_term n c depth) d)) :: replrec depth rest
+ Alias (initial,(na,b,map_pair (replace_term sigma n c depth) d)) :: replrec depth rest
| NonDepAlias :: rest ->
NonDepAlias :: replrec depth rest
| Abstract (i,d) :: rest ->
- Abstract (i, RelDecl.map_constr (replace_term n c depth) d)
+ Abstract (i, RelDecl.map_constr (fun t -> replace_term sigma n c depth t) d)
:: replrec (depth+1) rest in
replrec 0
@@ -702,7 +709,7 @@ let merge_name get_name obj = function
let merge_names get_name = List.map2 (merge_name get_name)
-let get_names env sign eqns =
+let get_names env sigma sign eqns =
let names1 = List.make (Context.Rel.length sign) Anonymous in
(* If any, we prefer names used in pats, from top to bottom *)
let names2,aliasname =
@@ -721,7 +728,7 @@ let get_names env sign eqns =
(fun (l,avoid) d na ->
let na =
merge_name
- (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env t na) avoid))
+ (fun (LocalAssum (na,t) | LocalDef (na,_,t)) -> Name (next_name_away (named_hd env sigma t na) avoid))
d na
in
(na::l,(out_name na)::avoid))
@@ -832,13 +839,13 @@ let rec map_predicate f k ccl = function
| Abstract _ :: rest ->
map_predicate f (k+1) ccl rest
-let noccur_predicate_between n = map_predicate (noccur_between n)
+let noccur_predicate_between sigma n = map_predicate (noccur_between sigma n)
let liftn_predicate n = map_predicate (liftn n)
let lift_predicate n = liftn_predicate n 1
-let regeneralize_index_predicate n = map_predicate (relocate_index n 1) 0
+let regeneralize_index_predicate sigma n = map_predicate (relocate_index sigma n 1) 0
let substnl_predicate sigma = map_predicate (substnl sigma)
@@ -859,6 +866,7 @@ let specialize_predicate_var (cur,typ,dep) env tms ccl =
| IsInd (_, IndType (_, _), []) -> []
| IsInd (_, IndType (indf, realargs), names) ->
let arsign,_ = get_arity env indf in
+ let arsign = List.map EConstr.of_rel_decl arsign in
subst_of_rel_context_instance arsign realargs
| NotInd _ -> [] in
subst_predicate (l,c) ccl tms
@@ -872,13 +880,13 @@ let specialize_predicate_var (cur,typ,dep) env tms ccl =
(* We first need to lift t(x) s.t. it is typed in Gamma, X:=rargs, x' *)
(* then we have to replace x by x' in t(x) and y by y' in P *)
(*****************************************************************************)
-let generalize_predicate (names,na) ny d tms ccl =
+let generalize_predicate sigma (names,na) ny d tms ccl =
let () = match na with
| Anonymous -> anomaly (Pp.str "Undetected dependency")
| _ -> () in
let p = List.length names + 1 in
let ccl = lift_predicate 1 ccl tms in
- regeneralize_index_predicate (ny+p+1) ccl tms
+ regeneralize_index_predicate sigma (ny+p+1) ccl tms
(*****************************************************************************)
(* We just matched over cur:ind(realargs) in the following matching problem *)
@@ -920,16 +928,16 @@ let rec extract_predicate ccl = function
ccl
let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
- let sign = make_arity_signature env true indf in
+ let sign = make_arity_signature env sigma true indf in
(* n is the number of real args + 1 (+ possible let-ins in sign) *)
let n = List.length sign in
(* Before abstracting we generalize over cur and on those realargs *)
(* that are rels, consistently with the specialization made in *)
(* build_branch *)
let tms = List.fold_right2 (fun par arg tomatch ->
- match kind_of_term par with
- | Rel i -> relocate_index_tomatch (i+n) (destRel arg) tomatch
- | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list 0 sign)
+ match EConstr.kind sigma par with
+ | Rel i -> relocate_index_tomatch sigma (i+n) (destRel sigma arg) tomatch
+ | _ -> tomatch) (realargs@[cur]) (Context.Rel.to_extended_list EConstr.mkRel 0 sign)
(lift_tomatch_stack n tms) in
(* Pred is already dependent in the current term to match (if *)
(* (na<>Anonymous) and its realargs; we just need to adjust it to *)
@@ -941,7 +949,7 @@ let abstract_predicate env sigma indf cur realargs (names,na) tms ccl =
let pred = extract_predicate ccl tms in
(* Build the predicate properly speaking *)
let sign = List.map2 set_name (na::names) sign in
- it_mkLambda_or_LetIn_name env pred sign
+ it_mkLambda_or_LetIn_name env sigma pred sign
(* [expand_arg] is used by [specialize_predicate]
if Yk denotes [Xk;xk] or [Xk],
@@ -976,11 +984,15 @@ let add_assert_false_case pb tomatch =
let adjust_impossible_cases pb pred tomatch submat =
match submat with
| [] ->
+ (** FIXME: This breaks if using evar-insensitive primitives. In particular,
+ 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
| 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
- pb.evdref := Evd.define evk default.uj_type evd
+ pb.evdref := Evd.define evk (EConstr.Unsafe.to_constr default.uj_type) evd
end;
add_assert_false_case pb tomatch
| _ ->
@@ -1026,12 +1038,13 @@ let specialize_predicate newtomatchs (names,depna) arsign cs tms ccl =
(* We prepare the substitution of X and x:I(X) *)
let realargsi =
if not (Int.equal nrealargs 0) then
- subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs)
+ CVars.subst_of_rel_context_instance arsign (Array.to_list cs.cs_concl_realargs)
else
[] in
+ let realargsi = List.map EConstr.of_constr realargsi in
let copti = match depna with
| Anonymous -> None
- | Name _ -> Some (build_dependent_constructor cs)
+ | Name _ -> Some (EConstr.of_constr (build_dependent_constructor cs))
in
(* The substituends realargsi, copti are all defined in gamma, x1...xn *)
(* We need _parallel_ bindings to get gamma, x1...xn |- PI tms. ccl'' *)
@@ -1067,69 +1080,69 @@ let adjust_predicate_from_tomatch tomatch (current,typ as ct) pb =
(* Remove commutative cuts that turn out to be non-dependent after
some evars have been instantiated *)
-let rec ungeneralize n ng body =
- match kind_of_term body with
+let rec ungeneralize sigma n ng body =
+ match EConstr.kind sigma body with
| Lambda (_,_,c) when Int.equal ng 0 ->
subst1 (mkRel n) c
| Lambda (na,t,c) ->
(* We traverse an inner generalization *)
- mkLambda (na,t,ungeneralize (n+1) (ng-1) c)
+ mkLambda (na,t,ungeneralize sigma (n+1) (ng-1) c)
| LetIn (na,b,t,c) ->
(* We traverse an alias *)
- mkLetIn (na,b,t,ungeneralize (n+1) ng c)
+ mkLetIn (na,b,t,ungeneralize sigma (n+1) ng c)
| Case (ci,p,c,brs) ->
(* We traverse a split *)
let p =
- let sign,p = decompose_lam_assum p in
- let sign2,p = decompose_prod_n_assum ng p in
- let p = prod_applist p [mkRel (n+List.length sign+ng)] in
+ let sign,p = decompose_lam_assum sigma p in
+ let sign2,p = decompose_prod_n_assum sigma ng p in
+ let p = prod_applist sigma p [mkRel (n+List.length sign+ng)] in
it_mkLambda_or_LetIn (it_mkProd_or_LetIn p sign2) sign in
mkCase (ci,p,c,Array.map2 (fun q c ->
- let sign,b = decompose_lam_n_decls q c in
- it_mkLambda_or_LetIn (ungeneralize (n+q) ng b) sign)
+ let sign,b = decompose_lam_n_decls sigma q c in
+ it_mkLambda_or_LetIn (ungeneralize sigma (n+q) ng b) sign)
ci.ci_cstr_ndecls brs)
| App (f,args) ->
(* We traverse an inner generalization *)
- assert (isCase f);
- mkApp (ungeneralize n (ng+Array.length args) f,args)
+ assert (isCase sigma f);
+ mkApp (ungeneralize sigma n (ng+Array.length args) f,args)
| _ -> assert false
-let ungeneralize_branch n k (sign,body) cs =
- (sign,ungeneralize (n+cs.cs_nargs) k body)
+let ungeneralize_branch sigma n k (sign,body) cs =
+ (sign,ungeneralize sigma (n+cs.cs_nargs) k body)
-let rec is_dependent_generalization ng body =
- match kind_of_term body with
+let rec is_dependent_generalization sigma ng body =
+ match EConstr.kind sigma body with
| Lambda (_,_,c) when Int.equal ng 0 ->
- dependent (mkRel 1) c
+ not (noccurn sigma 1 c)
| Lambda (na,t,c) ->
(* We traverse an inner generalization *)
- is_dependent_generalization (ng-1) c
+ is_dependent_generalization sigma (ng-1) c
| LetIn (na,b,t,c) ->
(* We traverse an alias *)
- is_dependent_generalization ng c
+ is_dependent_generalization sigma ng c
| Case (ci,p,c,brs) ->
(* We traverse a split *)
Array.exists2 (fun q c ->
- let _,b = decompose_lam_n_decls q c in
- is_dependent_generalization ng b)
+ let _,b = decompose_lam_n_decls sigma q c in
+ is_dependent_generalization sigma ng b)
ci.ci_cstr_ndecls brs
| App (g,args) ->
(* We traverse an inner generalization *)
- assert (isCase g);
- is_dependent_generalization (ng+Array.length args) g
+ assert (isCase sigma g);
+ is_dependent_generalization sigma (ng+Array.length args) g
| _ -> assert false
-let is_dependent_branch k (_,br) =
- is_dependent_generalization k br
+let is_dependent_branch sigma k (_,br) =
+ is_dependent_generalization sigma k br
let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
let rec aux k brs tomatch pred tocheck deps = match deps, tomatch with
| [], _ -> brs,tomatch,pred,[]
| n::deps, Abstract (i,d) :: tomatch ->
- let d = map_constr (nf_evar evd) d in
+ let d = map_constr (fun c -> nf_evar evd c) d in
let is_d = match d with LocalAssum _ -> false | LocalDef _ -> true in
- if is_d || List.exists (fun c -> dependent_decl (lift k c) d) tocheck
- && Array.exists (is_dependent_branch k) brs then
+ if is_d || List.exists (fun c -> dependent_decl evd (lift k c) d) tocheck
+ && Array.exists (is_dependent_branch evd k) brs then
(* Dependency in the current term to match and its dependencies is real *)
let brs,tomatch,pred,inst = aux (k+1) brs tomatch pred (mkRel n::tocheck) deps in
let inst = match d with
@@ -1142,9 +1155,9 @@ let postprocess_dependencies evd tocheck brs tomatch pred deps cs =
(* terms by its actual value in both the remaining terms to match and *)
(* the bodies of the Case *)
let pred = lift_predicate (-1) pred tomatch in
- let tomatch = relocate_index_tomatch 1 (n+1) tomatch in
+ let tomatch = relocate_index_tomatch evd 1 (n+1) tomatch in
let tomatch = lift_tomatch_stack (-1) tomatch in
- let brs = Array.map2 (ungeneralize_branch n k) brs cs in
+ let brs = Array.map2 (ungeneralize_branch evd n k) brs cs in
aux k brs tomatch pred tocheck deps
| _ -> assert false
in aux 0 brs tomatch pred tocheck deps
@@ -1196,17 +1209,17 @@ let rec generalize_problem names pb = function
| [] -> pb, []
| i::l ->
let pb',deps = generalize_problem names pb l in
- let d = map_constr (lift i) (Environ.lookup_rel i pb.env) in
+ let d = map_constr (lift i) (lookup_rel i pb.env) in
begin match d with
| LocalDef (Anonymous,_,_) -> pb', deps
| _ ->
(* for better rendering *)
- let d = RelDecl.map_type (whd_betaiota !(pb.evdref)) d in
+ let d = RelDecl.map_type (fun c -> whd_betaiota !(pb.evdref) c) d in
let tomatch = lift_tomatch_stack 1 pb'.tomatch in
- let tomatch = relocate_index_tomatch (i+1) 1 tomatch in
+ let tomatch = relocate_index_tomatch !(pb.evdref) (i+1) 1 tomatch in
{ pb' with
tomatch = Abstract (i,d) :: tomatch;
- pred = generalize_predicate names i d pb'.tomatch pb'.pred },
+ pred = generalize_predicate !(pb'.evdref) names i d pb'.tomatch pb'.pred },
i::deps
end
@@ -1228,7 +1241,8 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
(* build the name x1..xn from the names present in the equations *)
(* that had matched constructor C *)
let cs_args = const_info.cs_args in
- let names,aliasname = get_names pb.env cs_args eqns in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs_args in
+ let names,aliasname = get_names pb.env !(pb.evdref) cs_args eqns in
let typs = List.map2 RelDecl.set_name names cs_args
in
@@ -1251,30 +1265,30 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
(* We compute over which of x(i+1)..xn and x matching on xi will need a *)
(* generalization *)
let dep_sign =
- find_dependencies_signature
- (dependencies_in_rhs const_info.cs_nargs current pb.tomatch eqns)
+ find_dependencies_signature !(pb.evdref)
+ (dependencies_in_rhs !(pb.evdref) const_info.cs_nargs current pb.tomatch eqns)
(List.rev typs') in
(* The dependent term to subst in the types of the remaining UnPushed
terms is relative to the current context enriched by topushs *)
- let ci = build_dependent_constructor const_info in
+ let ci = EConstr.of_constr (build_dependent_constructor const_info) in
(* Current context Gamma has the form Gamma1;cur:I(realargs);Gamma2 *)
(* We go from Gamma |- PI tms. pred to *)
(* Gamma;x1..xn;curalias:I(x1..xn) |- PI tms'. pred' *)
(* where, in tms and pred, those realargs that are vars are *)
(* replaced by the corresponding xi and cur replaced by curalias *)
- let cirealargs = Array.to_list const_info.cs_concl_realargs in
+ let cirealargs = Array.map_to_list EConstr.of_constr const_info.cs_concl_realargs in
(* Do the specialization for terms to match *)
let tomatch = List.fold_right2 (fun par arg tomatch ->
- match kind_of_term par with
- | Rel i -> replace_tomatch (i+const_info.cs_nargs) arg tomatch
+ match EConstr.kind !(pb.evdref) par with
+ | Rel i -> replace_tomatch !(pb.evdref) (i+const_info.cs_nargs) arg tomatch
| _ -> tomatch) (current::realargs) (ci::cirealargs)
(lift_tomatch_stack const_info.cs_nargs pb.tomatch) in
let pred_is_not_dep =
- noccur_predicate_between 1 (List.length realnames + 1) pb.pred tomatch in
+ noccur_predicate_between !(pb.evdref) 1 (List.length realnames + 1) pb.pred tomatch in
let typs' =
List.map2
@@ -1300,10 +1314,10 @@ let build_branch initial current realargs deps (realnames,curname) pb arsign eqn
| Name _ ->
let cur_alias = lift const_info.cs_nargs current in
let ind =
- appvect (
- applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), snd const_info.cs_cstr),
- List.map (lift const_info.cs_nargs) const_info.cs_params),
- const_info.cs_concl_realargs) in
+ mkApp (
+ applist (mkIndU (inductive_of_constructor (fst const_info.cs_cstr), EInstance.make (snd const_info.cs_cstr)),
+ List.map (EConstr.of_constr %> lift const_info.cs_nargs) const_info.cs_params),
+ Array.map EConstr.of_constr const_info.cs_concl_realargs) in
Alias (initial,(aliasname,cur_alias,(ci,ind))) in
let tomatch = List.rev_append (alias :: currents) tomatch in
@@ -1369,7 +1383,7 @@ and match_current pb (initial,tomatch) =
(* We compile branches *)
let brvals = Array.map2 (compile_branch initial current realargs (names,dep) deps pb arsign) eqns cstrs in
(* We build the (elementary) case analysis *)
- let depstocheck = current::binding_vars_of_inductive typ in
+ let depstocheck = current::binding_vars_of_inductive !(pb.evdref) typ in
let brvals,tomatch,pred,inst =
postprocess_dependencies !(pb.evdref) depstocheck
brvals pb.tomatch pb.pred deps cstrs in
@@ -1381,11 +1395,11 @@ and match_current pb (initial,tomatch) =
let ci = make_case_info pb.env (fst mind) pb.casestyle in
let pred = nf_betaiota !(pb.evdref) pred in
let case =
- make_case_or_project pb.env indf ci pred current brvals
+ make_case_or_project pb.env !(pb.evdref) indf ci pred current brvals
in
Typing.check_allowed_sort pb.env !(pb.evdref) mind current pred;
{ uj_val = applist (case, inst);
- uj_type = prod_applist typ inst }
+ uj_type = prod_applist !(pb.evdref) typ inst }
(* Building the sub-problem when all patterns are variables. Case
@@ -1453,8 +1467,9 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest =
history = pop_history_pattern pb.history;
mat = List.map (push_alias_eqn alias) pb.mat } in
let j = compile pb in
+ let sigma = !(pb.evdref) in
{ uj_val =
- if isRel c || isVar c || count_occurrences (mkRel 1) j.uj_val <= 1 then
+ if isRel sigma c || isVar sigma c || count_occurrences sigma (mkRel 1) j.uj_val <= 1 then
subst1 c j.uj_val
else
mkLetIn (na,c,t,j.uj_val);
@@ -1479,10 +1494,10 @@ and compile_alias initial pb (na,orig,(expanded,expanded_typ)) rest =
evaluation; the drawback is that it might duplicate the instances
of the term to match when the corresponding variable is
substituted by a non-evaluated expression *)
- if not (Flags.is_program_mode ()) && (isRel orig || isVar orig) then
+ if not (Flags.is_program_mode ()) && (isRel sigma orig || isVar sigma orig) then
(* Try to compile first using non expanded alias *)
try
- if initial then f orig (Retyping.get_type_of pb.env !(pb.evdref) orig)
+ if initial then f orig (Retyping.get_type_of pb.env sigma orig)
else just_pop ()
with e when precatchable_exception e ->
(* Try then to compile using expanded alias *)
@@ -1563,7 +1578,7 @@ let matx_of_eqns env eqns =
returning True never happens and any inhabited type can be put instead).
*)
-let adjust_to_extended_env_and_remove_deps env extenv subst t =
+let adjust_to_extended_env_and_remove_deps env extenv sigma subst t =
let n = Context.Rel.length (rel_context env) in
let n' = Context.Rel.length (rel_context extenv) in
(* We first remove the bindings that are dependently typed (they are
@@ -1581,11 +1596,11 @@ let adjust_to_extended_env_and_remove_deps env extenv subst t =
let (p, _, _) = lookup_rel_id x (rel_context extenv) in
let rec traverse_local_defs p =
match lookup_rel p extenv with
- | LocalDef (_,c,_) -> assert (isRel c); traverse_local_defs (p + destRel c)
+ | LocalDef (_,c,_) -> assert (isRel sigma c); traverse_local_defs (p + destRel sigma c)
| LocalAssum _ -> p in
let p = traverse_local_defs p in
let u = lift (n' - n) u in
- try Some (p, u, expand_vars_in_term extenv u)
+ try Some (p, u, expand_vars_in_term extenv sigma u)
(* pedrot: does this really happen to raise [Failure _]? *)
with Failure _ -> None in
let subst0 = List.map_filter map subst in
@@ -1616,17 +1631,17 @@ let rec list_assoc_in_triple x = function
let abstract_tycon loc env evdref subst tycon extenv t =
let t = nf_betaiota !evdref t in (* it helps in some cases to remove K-redex*)
- let src = match kind_of_term t with
+ let src = match EConstr.kind !evdref t with
| Evar (evk,_) -> (loc,Evar_kinds.SubEvar evk)
| _ -> (loc,Evar_kinds.CasesType true) in
- let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv subst t in
+ let subst0,t0 = adjust_to_extended_env_and_remove_deps env extenv !evdref subst t in
(* We traverse the type T of the original problem Xi looking for subterms
that match the non-constructor part of the constraints (this part
is in subst); these subterms are the "good" subterms and we replace them
by an evar that may depend (and only depend) on the corresponding
convertible subterms of the substitution *)
let rec aux (k,env,subst as x) t =
- let t = whd_evar !evdref t in match kind_of_term t with
+ match EConstr.kind !evdref t with
| Rel n when is_local_def (lookup_rel n env) -> t
| Evar ev ->
let ty = get_type_of env !evdref t in
@@ -1646,7 +1661,7 @@ let abstract_tycon loc env evdref subst tycon extenv t =
let good = List.filter (fun (_,u,_) -> is_conv_leq env !evdref t u) subst in
match good with
| [] ->
- map_constr_with_full_binders push_binder aux x t
+ map_constr_with_full_binders !evdref push_binder aux x t
| (_, _, u) :: _ -> (* u is in extenv *)
let vl = List.map pi1 good in
let ty =
@@ -1654,16 +1669,16 @@ let abstract_tycon loc env evdref subst tycon extenv t =
Evarutil.evd_comb1 (refresh_universes (Some false) env) evdref ty
in
let ty = lift (-k) (aux x ty) in
- let depvl = free_rels ty in
+ let depvl = free_rels !evdref ty in
let inst =
List.map_i
(fun i _ -> if Int.List.mem i vl then u else mkRel i) 1
(rel_context extenv) in
let rel_filter =
- List.map (fun a -> not (isRel a) || dependent a u
- || Int.Set.mem (destRel a) depvl) inst in
+ List.map (fun a -> not (isRel !evdref a) || dependent !evdref a u
+ || Int.Set.mem (destRel !evdref a) depvl) inst in
let named_filter =
- List.map (fun d -> dependent (mkVar (NamedDecl.get_id d)) u)
+ List.map (fun d -> local_occur_var !evdref (NamedDecl.get_id d) u)
(named_context extenv) in
let filter = Filter.make (rel_filter @ named_filter) in
let candidates = u :: List.map mkRel vl in
@@ -1703,13 +1718,13 @@ let build_tycon loc env tycon_env s subst tycon extenv evdref t =
let build_inversion_problem loc env sigma tms t =
let make_patvar t (subst,avoid) =
- let id = next_name_away (named_hd env t Anonymous) avoid in
+ let id = next_name_away (named_hd env sigma t Anonymous) avoid in
PatVar (Loc.ghost,Name id), ((id,t)::subst, id::avoid) in
let rec reveal_pattern t (subst,avoid as acc) =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Construct (cstr,u) -> PatCstr (Loc.ghost,cstr,[],Anonymous), acc
- | App (f,v) when isConstruct f ->
- let cstr,u = destConstruct f in
+ | App (f,v) when isConstruct sigma f ->
+ let cstr,u = destConstruct sigma f in
let n = constructor_nrealargs_env env cstr in
let l = List.lastn n (Array.to_list v) in
let l,acc = List.fold_map' reveal_pattern l acc in
@@ -1722,7 +1737,7 @@ let build_inversion_problem loc env sigma tms t =
let patl,acc = List.fold_map' reveal_pattern realargs acc in
let pat,acc = make_patvar t acc in
let indf' = lift_inductive_family n indf in
- let sign = make_arity_signature env true indf' in
+ let sign = make_arity_signature env sigma true indf' in
let patl = pat :: List.rev patl in
let patl,sign = recover_and_adjust_alias_names patl sign in
let p = List.length patl in
@@ -1755,7 +1770,7 @@ let build_inversion_problem loc env sigma tms t =
List.map (fun (c,d) -> (c,extract_inductive_data pb_env sigma d,d)) decls in
let decls = List.rev decls in
- let dep_sign = find_dependencies_signature (List.make n true) decls in
+ let dep_sign = find_dependencies_signature sigma (List.make n true) decls in
let sub_tms =
List.map2 (fun deps (tm, (tmtyp,_), decl) ->
@@ -1845,6 +1860,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
let ((ind,u),_) = dest_ind_family indf' in
let nrealargs_ctxt = inductive_nrealdecls_env env0 ind in
let arsign = fst (get_arity env0 indf') in
+ let arsign = List.map (fun d -> map_rel_decl EConstr.of_constr d) arsign in
let realnal =
match t with
| Some (loc,ind',realnal) ->
@@ -1854,7 +1870,7 @@ let extract_arity_signature ?(dolift=true) env0 tomatchl tmsign =
anomaly (Pp.str "Ill-formed 'in' clause in cases");
List.rev realnal
| None -> List.make nrealargs_ctxt Anonymous in
- LocalAssum (na, build_dependent_inductive env0 indf')
+ LocalAssum (na, EConstr.of_constr (build_dependent_inductive env0 indf'))
::(List.map2 RelDecl.set_name realnal arsign) in
let rec buildrec n = function
| [],[] -> []
@@ -1879,8 +1895,8 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
let subst, len =
List.fold_right2 (fun (tm, tmtype) sign (subst, len) ->
let signlen = List.length sign in
- match kind_of_term tm with
- | Rel n when dependent tm c
+ match EConstr.kind sigma tm with
+ | Rel n when dependent sigma tm c
&& Int.equal signlen 1 (* The term to match is not of a dependent type itself *) ->
((n, len) :: subst, len - signlen)
| Rel n when signlen > 1 (* The term is of a dependent type,
@@ -1891,21 +1907,21 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
let subst, len =
List.fold_left
(fun (subst, len) arg ->
- match kind_of_term arg with
- | Rel n when dependent arg c ->
+ match EConstr.kind sigma arg with
+ | Rel n when dependent sigma arg c ->
((n, len) :: subst, pred len)
| _ -> (subst, pred len))
(subst, len) realargs
in
let subst =
- if dependent tm c && List.for_all isRel realargs
+ if dependent sigma tm c && List.for_all (isRel sigma) realargs
then (n, len) :: subst else subst
in (subst, pred len))
| _ -> (subst, len - signlen))
(List.rev tomatchs) arsign ([], nar)
in
let rec predicate lift c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Rel n when n > lift ->
(try
(* Make the predicate dependent on the matched variable *)
@@ -1915,7 +1931,7 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
(* A variable that is not matched, lift over the arsign. *)
mkRel (n + nar))
| _ ->
- map_constr_with_binders succ predicate lift c
+ EConstr.map_with_binders sigma succ predicate lift c
in
assert (len == 0);
let p = predicate 0 c in
@@ -1935,6 +1951,21 @@ let prepare_predicate_from_arsign_tycon env sigma loc tomatchs arsign c =
* tycon to make the predicate if it is not closed.
*)
+exception LocalOccur
+
+let noccur_with_meta sigma n m term =
+ let rec occur_rec n c = match EConstr.kind sigma c with
+ | Rel p -> if n<=p && p<n+m then raise LocalOccur
+ | App(f,cl) ->
+ (match EConstr.kind sigma f with
+ | Cast (c,_,_) when isMeta sigma c -> ()
+ | Meta _ -> ()
+ | _ -> EConstr.iter_with_binders sigma succ occur_rec n c)
+ | Evar (_, _) -> ()
+ | _ -> EConstr.iter_with_binders sigma succ occur_rec n c
+ in
+ try (occur_rec n term; true) with LocalOccur -> false
+
let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
let refresh_tycon sigma t =
(** If we put the typing constraint in the term, it has to be
@@ -1946,7 +1977,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
let preds =
match pred, tycon with
(* No return clause *)
- | None, Some t when not (noccur_with_meta 0 max_int t) ->
+ | None, Some t when not (noccur_with_meta sigma 0 max_int t) ->
(* If the tycon is not closed w.r.t real variables, we try *)
(* two different strategies *)
(* First strategy: we abstract the tycon wrt to the dependencies *)
@@ -1983,7 +2014,7 @@ let prepare_predicate loc typing_fun env sigma tomatchs arsign tycon pred =
let evdref = ref sigma in
let predcclj = typing_fun (mk_tycon (mkSort newt)) envar evdref rtntyp in
let sigma = !evdref in
- let predccl = (j_nf_evar sigma predcclj).uj_val in
+ let predccl = nf_evar sigma predcclj.uj_val in
[sigma, predccl]
in
List.map
@@ -2052,6 +2083,7 @@ let constr_of_pat env evdref arsign pat avoid =
{uj_val = ty; uj_type = Typing.unsafe_type_of env !evdref ty}
in
let (ind,u), params = dest_ind_family indf in
+ let params = List.map EConstr.of_constr params in
if not (eq_ind ind cind) then error_bad_constructor ~loc:l env cstr ind;
let cstrs = get_constructors env indf in
let ci = cstrs.(i-1) in
@@ -2060,7 +2092,7 @@ let constr_of_pat env evdref arsign pat avoid =
let patargs, args, sign, env, n, m, avoid =
List.fold_right2
(fun decl ua (patargs, args, sign, env, n, m, avoid) ->
- let t = RelDecl.get_type decl in
+ let t = EConstr.of_constr (RelDecl.get_type decl) in
let pat', sign', arg', typ', argtypargs, n', avoid =
let liftt = liftn (List.length sign) (succ (List.length args)) t in
typ env (substl args liftt, []) ua avoid
@@ -2073,11 +2105,11 @@ let constr_of_pat env evdref arsign pat avoid =
let args = List.rev args in
let patargs = List.rev patargs in
let pat' = PatCstr (l, cstr, patargs, alias) in
- let cstr = mkConstructU ci.cs_cstr in
- let app = applistc cstr (List.map (lift (List.length sign)) params) in
- let app = applistc app args in
+ let cstr = mkConstructU (on_snd EInstance.make ci.cs_cstr) in
+ let app = applist (cstr, List.map (lift (List.length sign)) params) in
+ let app = applist (app, args) in
let apptype = Retyping.get_type_of env ( !evdref) app in
- let IndType (indf, realargs) = find_rectype env ( !evdref) apptype in
+ let IndType (indf, realargs) = find_rectype env (!evdref) apptype in
match alias with
Anonymous ->
pat', sign, app, apptype, realargs, n, avoid
@@ -2111,22 +2143,22 @@ let eq_id avoid id =
avoid := hid' :: !avoid;
hid'
-let is_topvar t =
-match kind_of_term t with
+let is_topvar sigma t =
+match EConstr.kind sigma t with
| Rel 0 -> true
| _ -> false
-let rels_of_patsign =
+let rels_of_patsign sigma =
List.map (fun decl ->
match decl with
- | LocalDef (na,t',t) when is_topvar t' -> LocalAssum (na,t)
+ | LocalDef (na,t',t) when is_topvar sigma t' -> LocalAssum (na,t)
| _ -> decl)
-let vars_of_ctx ctx =
+let vars_of_ctx sigma ctx =
let _, y =
List.fold_right (fun decl (prev, vars) ->
match decl with
- | LocalDef (na,t',t) when is_topvar t' ->
+ | LocalDef (na,t',t) when is_topvar sigma t' ->
prev,
(GApp (Loc.ghost,
(GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)),
@@ -2146,6 +2178,9 @@ let rec is_included x y =
if Int.equal i i' then List.for_all2 is_included args args'
else false
+let lift_rel_context n l =
+ map_rel_context_with_binders (liftn n) l
+
(* liftsign is the current pattern's complete signature length.
Hence pats is already typed in its
full signature. However prevpatterns are in the original one signature per pattern form.
@@ -2220,12 +2255,12 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
(* lift to get outside of past patterns to get terms in the combined environment. *)
(fun (pats, n) (sign, c, (s, args), p) ->
let len = List.length sign in
- ((rels_of_patsign sign, lift n c,
+ ((rels_of_patsign !evdref sign, lift n c,
(s, List.map (lift n) args), p) :: pats, len + n))
([], 0) pats
in
let ineqs = build_ineqs evdref prevpatterns pats signlen in
- let rhs_rels' = rels_of_patsign rhs_rels in
+ let rhs_rels' = rels_of_patsign !evdref rhs_rels in
let _signenv = push_rel_context rhs_rels' env in
let arity =
let args, nargs =
@@ -2243,7 +2278,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
| Some ineqs ->
[LocalAssum (Anonymous, ineqs)], lift 1 arity
in
- let eqs_rels, arity = decompose_prod_n_assum neqs arity in
+ let eqs_rels, arity = decompose_prod_n_assum !evdref neqs arity in
eqs_rels @ neqs_rels @ rhs_rels', arity
in
let rhs_env = push_rel_context rhs_rels' env in
@@ -2255,7 +2290,7 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in
let branch =
let bref = GVar (Loc.ghost, branch_name) in
- match vars_of_ctx rhs_rels with
+ match vars_of_ctx !evdref rhs_rels with
[] -> bref
| l -> GApp (Loc.ghost, bref, l)
in
@@ -2289,16 +2324,16 @@ let lift_ctx n ctx =
in ctx'
(* Turn matched terms into variables. *)
-let abstract_tomatch env tomatchs tycon =
+let abstract_tomatch env sigma tomatchs tycon =
let prev, ctx, names, tycon =
List.fold_left
(fun (prev, ctx, names, tycon) (c, t) ->
let lenctx = List.length ctx in
- match kind_of_term c with
+ match EConstr.kind sigma c with
Rel n -> (lift lenctx c, lift_tomatch_type lenctx t) :: prev, ctx, names, tycon
| _ ->
let tycon = Option.map
- (fun t -> subst_term (lift 1 c) (lift 1 t)) tycon in
+ (fun t -> subst_term sigma (lift 1 c) (lift 1 t)) tycon in
let name = next_ident_away (Id.of_string "filtered_var") names in
(mkRel 1, lift_tomatch_type (succ lenctx) t) :: lift_ctx 1 prev,
LocalDef (Name name, lift lenctx c, lift lenctx $ type_of_tomatch t) :: ctx,
@@ -2350,7 +2385,7 @@ let build_dependent_signature env evdref avoid tomatchs arsign =
in
let previd, id =
let name =
- match kind_of_term arg with
+ match EConstr.kind !evdref arg with
Rel n -> RelDecl.get_name (lookup_rel n env)
| _ -> name
in
@@ -2416,7 +2451,7 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
(* constructors found in patterns *)
let tomatchs = coerce_to_indtype typing_function evdref env matx tomatchl in
let tycon = valcon_of_tycon tycon in
- let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env tomatchs tycon in
+ let tomatchs, tomatchs_lets, tycon' = abstract_tomatch env !evdref tomatchs tycon in
let env = push_rel_context tomatchs_lets env in
let len = List.length eqns in
let sign, allnames, signlen, eqs, neqs, args =
@@ -2470,14 +2505,14 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
List.map (fun (c,d) -> (c,extract_inductive_data env !evdref d,d)) typs in
let dep_sign =
- find_dependencies_signature
+ find_dependencies_signature !evdref
(List.make (List.length typs) true)
typs in
let typs' =
List.map3
(fun (tm,tmt) deps na ->
- let deps = if not (isRel tm) then [] else deps in
+ let deps = if not (isRel !evdref tm) then [] else deps in
((tm,tmt),deps,na))
tomatchs dep_sign nal in
@@ -2501,10 +2536,10 @@ let compile_program_cases loc style (typing_function, evdref) tycon env
let j = compile pb in
(* We check for unused patterns *)
List.iter (check_unused_pattern env) matx;
- let body = it_mkLambda_or_LetIn (applistc j.uj_val args) lets in
+ let body = it_mkLambda_or_LetIn (applist (j.uj_val, args)) lets in
let j =
{ uj_val = it_mkLambda_or_LetIn body tomatchs_lets;
- uj_type = nf_evar !evdref tycon; }
+ uj_type = EConstr.of_constr (EConstr.to_constr !evdref tycon); }
in j
(**************************************************************************)
@@ -2545,14 +2580,14 @@ let compile_cases loc style (typing_fun, evdref) tycon env (predopt, tomatchl, e
List.map (fun (c,d) -> (c,extract_inductive_data env sigma d,d)) typs in
let dep_sign =
- find_dependencies_signature
+ find_dependencies_signature !evdref
(List.make (List.length typs) true)
typs in
let typs' =
List.map3
(fun (tm,tmt) deps na ->
- let deps = if not (isRel tm) then [] else deps in
+ let deps = if not (isRel !evdref tm) then [] else deps in
((tm,tmt),deps,na))
tomatchs dep_sign nal in
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index d8fad1687..6c2b5bf68 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -10,6 +10,7 @@ open Names
open Term
open Evd
open Environ
+open EConstr
open Inductiveops
open Glob_term
open Evarutil
@@ -46,12 +47,12 @@ val compile_cases :
val constr_of_pat :
Environ.env ->
Evd.evar_map ref ->
- Context.Rel.Declaration.t list ->
+ rel_context ->
Glob_term.cases_pattern ->
Names.Id.t list ->
Glob_term.cases_pattern *
- (Context.Rel.Declaration.t list * Term.constr *
- (Term.types * Term.constr list) * Glob_term.cases_pattern) *
+ (rel_context * constr *
+ (types * constr list) * Glob_term.cases_pattern) *
Names.Id.t list
type 'a rhs =
@@ -84,7 +85,7 @@ type tomatch_status =
| Pushed of (bool*((constr * tomatch_type) * int list * Name.t))
| Alias of (bool * (Name.t * constr * (constr * types)))
| NonDepAlias
- | Abstract of int * Context.Rel.Declaration.t
+ | Abstract of int * rel_declaration
type tomatch_stack = tomatch_status list
@@ -110,15 +111,14 @@ type 'a pattern_matching_problem =
typing_function: type_constraint -> env -> evar_map ref -> 'a option -> unsafe_judgment }
-val compile : 'a pattern_matching_problem -> Environ.unsafe_judgment
+val compile : 'a pattern_matching_problem -> unsafe_judgment
val prepare_predicate : Loc.t ->
(Evarutil.type_constraint ->
- Environ.env -> Evd.evar_map ref -> glob_constr -> Environ.unsafe_judgment) ->
+ Environ.env -> Evd.evar_map ref -> glob_constr -> unsafe_judgment) ->
Environ.env ->
Evd.evar_map ->
- (Term.types * tomatch_type) list ->
- Context.Rel.t list ->
- Constr.constr option ->
- glob_constr option ->
- (Evd.evar_map * Names.name list * Term.constr) list
+ (types * tomatch_type) list ->
+ rel_context list ->
+ constr option ->
+ glob_constr option -> (Evd.evar_map * Names.name list * constr) list
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 84bf849e7..e18625c42 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -376,7 +376,8 @@ and cbv_norm_value info = function (* reduction under binders *)
(* with profiling *)
let cbv_norm infos constr =
- with_stats (lazy (cbv_norm_term infos (subs_id 0) constr))
+ let constr = EConstr.Unsafe.to_constr constr in
+ EConstr.of_constr (with_stats (lazy (cbv_norm_term infos (subs_id 0) constr)))
type cbv_infos = cbv_value infos
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index 87a03abbd..b014af2c7 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Environ
open CClosure
open Esubst
@@ -23,6 +24,9 @@ val cbv_norm : cbv_infos -> constr -> constr
(***********************************************************************
i This is for cbv debug *)
+
+open Term
+
type cbv_value =
| VAL of int * constr
| STACK of int * cbv_value * cbv_stack
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 30d100af9..632ba0d9c 100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -51,6 +51,7 @@ type coe_info_typ = {
coe_param : int }
let coe_info_typ_equal c1 c2 =
+ let eq_constr c1 c2 = Termops.eq_constr Evd.empty (EConstr.of_constr c1) (EConstr.of_constr c2) in
eq_constr c1.coe_value c2.coe_value &&
eq_constr c1.coe_type c2.coe_type &&
c1.coe_local == c2.coe_local &&
@@ -192,15 +193,16 @@ let coercion_exists coe = CoeTypMap.mem coe !coercion_tab
(* find_class_type : evar_map -> constr -> cl_typ * universe_list * constr list *)
let find_class_type sigma t =
+ let open EConstr in
let t', args = Reductionops.whd_betaiotazeta_stack sigma t in
- match kind_of_term t' with
- | Var id -> CL_SECVAR id, Univ.Instance.empty, args
+ match EConstr.kind sigma t' with
+ | Var id -> CL_SECVAR id, EInstance.empty, args
| Const (sp,u) -> CL_CONST sp, u, args
| Proj (p, c) when not (Projection.unfolded p) ->
- CL_PROJ (Projection.constant p), Univ.Instance.empty, c :: args
+ CL_PROJ (Projection.constant p), EInstance.empty, (c :: args)
| Ind (ind_sp,u) -> CL_IND ind_sp, u, args
- | Prod (_,_,_) -> CL_FUN, Univ.Instance.empty, []
- | Sort _ -> CL_SORT, Univ.Instance.empty, []
+ | Prod (_,_,_) -> CL_FUN, EInstance.empty, []
+ | Sort _ -> CL_SORT, EInstance.empty, []
| _ -> raise Not_found
@@ -214,7 +216,7 @@ let subst_cl_typ subst ct = match ct with
| CL_CONST c ->
let c',t = subst_con_kn subst c in
if c' == c then ct else
- pi1 (find_class_type Evd.empty t)
+ pi1 (find_class_type Evd.empty (EConstr.of_constr t))
| CL_IND i ->
let i' = subst_ind subst i in
if i' == i then ct else CL_IND i'
@@ -297,9 +299,9 @@ let lookup_path_to_sort_from env sigma s =
let get_coercion_constructor env coe =
let c, _ =
- Reductionops.whd_all_stack env Evd.empty coe.coe_value
+ Reductionops.whd_all_stack env Evd.empty (EConstr.of_constr coe.coe_value)
in
- match kind_of_term c with
+ match EConstr.kind Evd.empty (** FIXME *) c with
| Construct (cstr,u) ->
(cstr, Inductiveops.constructor_nrealargs cstr -1)
| _ ->
@@ -317,7 +319,7 @@ let coercion_value { coe_value = c; coe_type = t; coe_context = ctx;
let subst, ctx = Universes.fresh_universe_context_set_instance ctx in
let c' = Vars.subst_univs_level_constr subst c
and t' = Vars.subst_univs_level_constr subst t in
- (make_judge c' t', b, b'), ctx
+ (make_judge (EConstr.of_constr c') (EConstr.of_constr t'), b, b'), ctx
(* pretty-print functions are now in Pretty *)
(* rajouter une coercion dans le graphe *)
@@ -403,7 +405,7 @@ type coercion = {
let reference_arity_length ref =
let t = Universes.unsafe_type_of_global ref in
- List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty t))
+ List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty (EConstr.of_constr t))) (** FIXME *)
let projection_arity_length p =
let len = reference_arity_length (ConstRef p) in
@@ -440,7 +442,8 @@ let cache_coercion (_, c) =
let is, _ = class_info c.coercion_source in
let it, _ = class_info c.coercion_target in
let value, ctx = Universes.fresh_global_instance (Global.env()) c.coercion_type in
- let typ = Retyping.get_type_of (Global.env ()) Evd.empty value in
+ let typ = Retyping.get_type_of (Global.env ()) Evd.empty (EConstr.of_constr value) in
+ let typ = EConstr.Unsafe.to_constr typ in
let xf =
{ coe_value = value;
coe_type = typ;
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index d509739cf..0d741a5a5 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -8,8 +8,9 @@
open Names
open Term
-open Evd
open Environ
+open EConstr
+open Evd
open Mod_subst
(** {6 This is the type of class kinds } *)
@@ -59,7 +60,7 @@ val class_info_from_index : cl_index -> cl_typ * cl_info_typ
(** [find_class_type env sigma c] returns the head reference of [c],
its universe instance and its arguments *)
-val find_class_type : evar_map -> types -> cl_typ * Univ.universe_instance * constr list
+val find_class_type : evar_map -> types -> cl_typ * EInstance.t * constr list
(** raises [Not_found] if not convertible to a class *)
val class_of : env -> evar_map -> types -> types * cl_index
@@ -104,7 +105,7 @@ val install_path_printer :
val string_of_class : cl_typ -> string
val pr_class : cl_typ -> std_ppcmds
val pr_cl_index : cl_index -> std_ppcmds
-val get_coercion_value : coe_index -> constr
+val get_coercion_value : coe_index -> Constr.t
val inheritance_graph : unit -> ((cl_index * cl_index) * inheritance_path) list
val classes : unit -> cl_typ list
val coercions : unit -> coe_index list
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index 2b860ae9c..542db7fdf 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -18,9 +18,10 @@ open CErrors
open Util
open Names
open Term
+open Environ
+open EConstr
open Vars
open Reductionops
-open Environ
open Typeops
open Pretype_errors
open Classops
@@ -52,7 +53,7 @@ let apply_coercion_args env evd check isproj argl funj =
let rec apply_rec acc typ = function
| [] ->
if isproj then
- let cst = fst (destConst (j_val funj)) in
+ let cst = fst (destConst !evdref (j_val funj)) in
let p = Projection.make cst false in
let pb = lookup_projection p env in
let args = List.skipn pb.Declarations.proj_npars argl in
@@ -63,9 +64,9 @@ let apply_coercion_args env evd check isproj argl funj =
{ uj_val = applist (j_val funj,argl);
uj_type = typ }
| h::restl -> (* On devrait pouvoir s'arranger pour qu'on n'ait pas a faire hnf_constr *)
- match kind_of_term (whd_all env evd typ) with
+ match EConstr.kind !evdref (whd_all env !evdref typ) with
| Prod (_,c1,c2) ->
- if check && not (e_cumul env evdref (Retyping.get_type_of env evd h) c1) then
+ if check && not (e_cumul env evdref (Retyping.get_type_of env !evdref h) c1) then
raise NoCoercion;
apply_rec (h::acc) (subst1 h c2) restl
| _ -> anomaly (Pp.str "apply_coercion_args")
@@ -99,10 +100,10 @@ let app_opt env evdref f t =
let pair_of_array a = (a.(0), a.(1))
-let disc_subset x =
- match kind_of_term x with
+let disc_subset sigma x =
+ match EConstr.kind sigma x with
| App (c, l) ->
- (match kind_of_term c with
+ (match EConstr.kind sigma c with
Ind (i,_) ->
let len = Array.length l in
let sigty = delayed_force sig_typ in
@@ -129,7 +130,7 @@ let lift_args n sign =
let mu env evdref t =
let rec aux v =
let v' = hnf env !evdref v in
- match disc_subset v' with
+ match disc_subset !evdref v' with
| Some (u, p) ->
let f, ct = aux u in
let p = hnf_nodelta env !evdref p in
@@ -140,8 +141,8 @@ let mu env evdref t =
| None -> (None, v)
in aux t
-and coerce loc env evdref (x : Term.constr) (y : Term.constr)
- : (Term.constr -> Term.constr) option
+and coerce loc env evdref (x : EConstr.constr) (y : EConstr.constr)
+ : (EConstr.constr -> EConstr.constr) option
=
let open Context.Rel.Declaration in
let rec coerce_unify env x y =
@@ -150,11 +151,11 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
evdref := the_conv_x_leq env x y !evdref;
None
with UnableToUnify _ -> coerce' env x y
- and coerce' env x y : (Term.constr -> Term.constr) option =
+ and coerce' env x y : (EConstr.constr -> EConstr.constr) option =
let subco () = subset_coerce env evdref x y in
let dest_prod c =
- match Reductionops.splay_prod_n env ( !evdref) 1 c with
- | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na,t), c
+ match Reductionops.splay_prod_n env (!evdref) 1 c with
+ | [LocalAssum (na,t) | LocalDef (na,_,t)], c -> (na, t), c
| _ -> raise NoSubtacCoercion
in
let coerce_application typ typ' c c' l l' =
@@ -174,12 +175,12 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
with UnableToUnify _ -> raise NoSubtacCoercion
in
(* Disallow equalities on arities *)
- if Reduction.is_arity env eqT then raise NoSubtacCoercion;
+ if Reductionops.is_arity env !evdref eqT then raise NoSubtacCoercion;
let restargs = lift_args 1
(List.rev (Array.to_list (Array.sub l (succ i) (len - (succ i)))))
in
let args = List.rev (restargs @ mkRel 1 :: List.map (lift 1) tele) in
- let pred = mkLambda (n, eqT, applistc (lift 1 c) args) in
+ let pred = mkLambda (n, eqT, applist (lift 1 c, args)) in
let eq = papp evdref coq_eq_ind [| eqT; hdx; hdy |] in
let evar = make_existential loc env evdref eq in
let eq_app x = papp evdref coq_eq_rect
@@ -191,14 +192,14 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
let term = co x in
Typing.e_solve_evars env evdref term)
in
- if isEvar c || isEvar c' || not (Program.is_program_generalized_coercion ()) then
+ if isEvar !evdref c || isEvar !evdref c' || not (Program.is_program_generalized_coercion ()) then
(* Second-order unification needed. *)
raise NoSubtacCoercion;
aux [] typ typ' 0 (fun x -> x)
in
- match (kind_of_term x, kind_of_term y) with
+ match (EConstr.kind !evdref x, EConstr.kind !evdref y) with
| Sort s, Sort s' ->
- (match s, s' with
+ (match ESorts.kind !evdref s, ESorts.kind !evdref s' with
| Prop x, Prop y when x == y -> None
| Prop _, Type _ -> None
| Type x, Type y when Univ.Universe.equal x y -> None (* false *)
@@ -224,7 +225,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
(mkApp (lift 1 f, [| coec1 |])))))
| App (c, l), App (c', l') ->
- (match kind_of_term c, kind_of_term c' with
+ (match EConstr.kind !evdref c, EConstr.kind !evdref c' with
Ind (i, u), Ind (i', u') -> (* Inductive types *)
let len = Array.length l in
let sigT = delayed_force sigT_typ in
@@ -241,16 +242,15 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
in
let c1 = coerce_unify env a a' in
let remove_head a c =
- match kind_of_term c with
+ match EConstr.kind !evdref c with
| Lambda (n, t, t') -> c, t'
| Evar (k, args) ->
let (evs, t) = Evardefine.define_evar_as_lambda env !evdref (k,args) in
evdref := evs;
- let (n, dom, rng) = destLambda t in
- let dom = whd_evar !evdref dom in
- if isEvar dom then
- let (domk, args) = destEvar dom in
- evdref := define domk a !evdref;
+ let (n, dom, rng) = destLambda !evdref t in
+ if isEvar !evdref dom then
+ let (domk, args) = destEvar !evdref dom in
+ evdref := define domk (EConstr.Unsafe.to_constr a) !evdref;
else ();
t, rng
| _ -> raise NoSubtacCoercion
@@ -301,7 +301,7 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
coerce_application typ typ' c c' l l')
else
subco ()
- | x, y when Constr.equal c c' ->
+ | x, y when EConstr.eq_constr !evdref c c' ->
if Int.equal (Array.length l) (Array.length l') then
let evm = !evdref in
let lam_type = Typing.unsafe_type_of env evm c in
@@ -312,14 +312,14 @@ and coerce loc env evdref (x : Term.constr) (y : Term.constr)
| _, _ -> subco ()
and subset_coerce env evdref x y =
- match disc_subset x with
+ match disc_subset !evdref x with
Some (u, p) ->
let c = coerce_unify env u y in
let f x =
app_opt env evdref c (papp evdref sig_proj1 [| u; p; x |])
in Some f
| None ->
- match disc_subset y with
+ match disc_subset !evdref y with
Some (u, p) ->
let c = coerce_unify env x u in
Some
@@ -374,7 +374,7 @@ let apply_coercion env sigma p hj typ_cl =
(* Try to coerce to a funclass; raise NoCoercion if not possible *)
let inh_app_fun_core env evd j =
let t = whd_all env evd j.uj_type in
- match kind_of_term t with
+ match EConstr.kind evd t with
| Prod (_,_,_) -> (evd,j)
| Evar ev ->
let (evd',t) = Evardefine.define_evar_as_product evd ev in
@@ -404,20 +404,25 @@ let inh_app_fun resolve_tc env evd j =
try inh_app_fun_core env (saturate_evd env evd) j
with NoCoercion -> (evd, j)
+let type_judgment env sigma j =
+ match EConstr.kind sigma (whd_all env sigma j.uj_type) with
+ | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind sigma s }
+ | _ -> error_not_a_type env sigma j
+
let inh_tosort_force loc env evd j =
try
let t,p = lookup_path_to_sort_from env evd j.uj_type in
let evd,j1 = apply_coercion env evd p j t in
let j2 = on_judgment_type (whd_evar evd) j1 in
- (evd,type_judgment env j2)
+ (evd,type_judgment env evd j2)
with Not_found | NoCoercion ->
error_not_a_type ~loc env evd j
let inh_coerce_to_sort loc env evd j =
let typ = whd_all env evd j.uj_type in
- match kind_of_term typ with
- | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = s })
- | Evar ev when not (is_defined evd (fst ev)) ->
+ match EConstr.kind evd typ with
+ | Sort s -> (evd,{ utj_val = j.uj_val; utj_type = ESorts.kind evd s })
+ | Evar ev ->
let (evd',s) = Evardefine.define_evar_as_sort env evd ev in
(evd',{ utj_val = j.uj_val; utj_type = s })
| _ ->
@@ -428,7 +433,7 @@ let inh_coerce_to_base loc env evd j =
let evdref = ref evd in
let ct, typ' = mu env evdref j.uj_type in
let res =
- { uj_val = app_coercion env evdref ct j.uj_val;
+ { uj_val = (app_coercion env evdref ct j.uj_val);
uj_type = typ' }
in !evdref, res
else (evd, j)
@@ -441,7 +446,7 @@ let inh_coerce_to_prod loc env evd t =
else (evd, t)
let inh_coerce_to_fail env evd rigidonly v t c1 =
- if rigidonly && not (Heads.is_rigid env c1 && Heads.is_rigid env t)
+ if rigidonly && not (Heads.is_rigid env (EConstr.Unsafe.to_constr c1) && Heads.is_rigid env (EConstr.Unsafe.to_constr t))
then
raise NoCoercion
else
@@ -466,8 +471,8 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 =
try inh_coerce_to_fail env evd rigidonly v t c1
with NoCoercion ->
match
- kind_of_term (whd_all env evd t),
- kind_of_term (whd_all env evd c1)
+ EConstr.kind evd (whd_all env evd t),
+ EConstr.kind evd (whd_all env evd c1)
with
| Prod (name,t1,t2), Prod (_,u1,u2) ->
(* Conversion did not work, we may succeed with a coercion. *)
@@ -485,9 +490,9 @@ let rec inh_conv_coerce_to_fail loc env evd rigidonly v t c1 =
inh_conv_coerce_to_fail loc env1 evd rigidonly
(Some (mkRel 1)) (lift 1 u1) (lift 1 t1) in
let v1 = Option.get v1 in
- let v2 = Option.map (fun v -> beta_applist (lift 1 v,[v1])) v in
+ let v2 = Option.map (fun v -> beta_applist evd' (lift 1 v,[v1])) v in
let t2 = match v2 with
- | None -> subst_term v1 t2
+ | None -> subst_term evd' v1 t2
| Some v2 -> Retyping.get_type_of env1 evd' v2 in
let (evd'',v2') = inh_conv_coerce_to_fail loc env1 evd' rigidonly v2 t2 u2 in
(evd'', Option.map (fun v2' -> mkLambda (name, u1, v2')) v2')
diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli
index 68f9a2e68..bc63d092d 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -10,6 +10,7 @@ open Evd
open Names
open Term
open Environ
+open EConstr
open Glob_term
(** {6 Coercions. } *)
diff --git a/pretyping/constr_matching.ml b/pretyping/constr_matching.ml
index 1cae8d16d..efe03bc2e 100644
--- a/pretyping/constr_matching.ml
+++ b/pretyping/constr_matching.ml
@@ -16,6 +16,7 @@ open Nameops
open Termops
open Reductionops
open Term
+open EConstr
open Vars
open Pattern
open Patternops
@@ -57,14 +58,15 @@ let warn_meta_collision =
strbrk " and a metavariable of same name.")
-let constrain n (ids, m as x) (names, terms as subst) =
+let constrain sigma n (ids, m) (names, terms as subst) =
+ let open EConstr in
try
let (ids', m') = Id.Map.find n terms in
- if List.equal Id.equal ids ids' && eq_constr m m' then subst
+ if List.equal Id.equal ids ids' && eq_constr sigma m m' then subst
else raise PatternMatchingFailure
with Not_found ->
let () = if Id.Map.mem n names then warn_meta_collision n in
- (names, Id.Map.add n x terms)
+ (names, Id.Map.add n (ids, m) terms)
let add_binders na1 na2 binding_vars (names, terms as subst) =
match na1, na2 with
@@ -79,11 +81,12 @@ let add_binders na1 na2 binding_vars (names, terms as subst) =
(names, terms)
| _ -> subst
-let rec build_lambda vars ctx m = match vars with
+let rec build_lambda sigma vars ctx m = match vars with
| [] ->
let len = List.length ctx in
- lift (-1 * len) m
+ EConstr.Vars.lift (-1 * len) m
| n :: vars ->
+ let open EConstr in
(* change [ x1 ... xn y z1 ... zm |- t ] into
[ x1 ... xn z1 ... zm |- lam y. t ] *)
let len = List.length ctx in
@@ -92,7 +95,7 @@ let rec build_lambda vars ctx m = match vars with
else if Int.equal i (pred n) then mkRel 1
else mkRel (i + 1)
in
- let m = substl (List.init len init) m in
+ let m = Vars.substl (List.init len init) m in
let pre, suf = List.chop (pred n) ctx in
match suf with
| [] -> assert false
@@ -100,12 +103,12 @@ let rec build_lambda vars ctx m = match vars with
let map i = if i > n then pred i else i in
let vars = List.map map vars in
(** Check that the abstraction is legal *)
- let frels = free_rels t in
+ let frels = free_rels sigma t in
let brels = List.fold_right Int.Set.add vars Int.Set.empty in
let () = if not (Int.Set.subset frels brels) then raise PatternMatchingFailure in
(** Create the abstraction *)
let m = mkLambda (na, t, m) in
- build_lambda vars (pre @ suf) m
+ build_lambda sigma vars (pre @ suf) m
let rec extract_bound_aux k accu frels ctx = match ctx with
| [] -> accu
@@ -123,41 +126,46 @@ let rec extract_bound_aux k accu frels ctx = match ctx with
let extract_bound_vars frels ctx =
extract_bound_aux 1 Id.Set.empty frels ctx
-let dummy_constr = mkProp
+let dummy_constr = EConstr.mkProp
let make_renaming ids = function
| (Name id, Name _, _) ->
begin
- try mkRel (List.index Id.equal id ids)
+ try EConstr.mkRel (List.index Id.equal id ids)
with Not_found -> dummy_constr
end
| _ -> dummy_constr
-let merge_binding allow_bound_rels ctx n cT subst =
+let to_fix (idx, (nas, cs, ts)) =
+ let inj = EConstr.of_constr in
+ (idx, (nas, Array.map inj cs, Array.map inj ts))
+
+let merge_binding sigma allow_bound_rels ctx n cT subst =
let c = match ctx with
| [] -> (* Optimization *)
([], cT)
| _ ->
- let frels = free_rels cT in
+ let frels = free_rels sigma cT in
if allow_bound_rels then
let vars = extract_bound_vars frels ctx in
let ordered_vars = Id.Set.elements vars in
let rename binding = make_renaming ordered_vars binding in
let renaming = List.map rename ctx in
- (ordered_vars, substl renaming cT)
+ (ordered_vars, Vars.substl renaming cT)
else
let depth = List.length ctx in
let min_elt = try Int.Set.min_elt frels with Not_found -> succ depth in
if depth < min_elt then
- ([], lift (- depth) cT)
+ ([], Vars.lift (- depth) cT)
else raise PatternMatchingFailure
in
- constrain n c subst
+ constrain sigma n c subst
let matches_core env sigma convert allow_partial_app allow_bound_rels
(binding_vars,pat) c =
+ let open EConstr in
let convref ref c =
- match ref, kind_of_term c with
+ match ref, EConstr.kind sigma c with
| VarRef id, Var id' -> Names.id_eq id id'
| ConstRef c, Const (c',_) -> Names.eq_constant c c'
| IndRef i, Ind (i', _) -> Names.eq_ind i i'
@@ -165,12 +173,12 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| _, _ ->
(if convert then
let sigma,c' = Evd.fresh_global env sigma ref in
- is_conv env sigma c' c
+ is_conv env sigma (EConstr.of_constr c') c
else false)
in
let rec sorec ctx env subst p t =
- let cT = strip_outer_cast t in
- match p,kind_of_term cT with
+ let cT = strip_outer_cast sigma t in
+ match p, EConstr.kind sigma cT with
| PSoApp (n,args),m ->
let fold (ans, seen) = function
| PRel n ->
@@ -179,13 +187,13 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| _ -> error "Only bound indices allowed in second order pattern matching."
in
let relargs, relset = List.fold_left fold ([], Int.Set.empty) args in
- let frels = free_rels cT in
+ let frels = free_rels sigma cT in
if Int.Set.subset frels relset then
- constrain n ([], build_lambda relargs ctx cT) subst
+ constrain sigma n ([], build_lambda sigma relargs ctx cT) subst
else
raise PatternMatchingFailure
- | PMeta (Some n), m -> merge_binding allow_bound_rels ctx n cT subst
+ | PMeta (Some n), m -> merge_binding sigma allow_bound_rels ctx n cT subst
| PMeta None, m -> subst
@@ -197,11 +205,14 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| PRel n1, Rel n2 when Int.equal n1 n2 -> subst
- | PSort GProp, Sort (Prop Null) -> subst
-
- | PSort GSet, Sort (Prop Pos) -> subst
+ | PSort ps, Sort s ->
- | PSort (GType _), Sort (Type _) -> subst
+ begin match ps, ESorts.kind sigma s with
+ | GProp, Prop Null -> subst
+ | GSet, Prop Pos -> subst
+ | GType _, Type _ -> subst
+ | _ -> raise PatternMatchingFailure
+ end
| PApp (p, [||]), _ -> sorec ctx env subst p t
@@ -216,10 +227,10 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
let subst =
match meta with
| None -> subst
- | Some n -> merge_binding allow_bound_rels ctx n c subst in
+ | Some n -> merge_binding sigma allow_bound_rels ctx n c subst in
Array.fold_left2 (sorec ctx env) subst args1 args22
else (* Might be a projection on the right *)
- match kind_of_term c2 with
+ match EConstr.kind sigma c2 with
| Proj (pr, c) when not (Projection.unfolded pr) ->
(try let term = Retyping.expand_projection env sigma pr c (Array.to_list args2) in
sorec ctx env subst p term
@@ -227,7 +238,7 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
| _ -> raise PatternMatchingFailure)
| PApp (c1,arg1), App (c2,arg2) ->
- (match c1, kind_of_term c2 with
+ (match c1, EConstr.kind sigma c2 with
| PRef (ConstRef r), Proj (pr,c) when not (eq_constant r (Projection.constant pr))
|| Projection.unfolded pr ->
raise PatternMatchingFailure
@@ -257,33 +268,33 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
sorec ctx env subst c1 c2
| PProd (na1,c1,d1), Prod(na2,c2,d2) ->
- sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env)
+ sorec ((na1,na2,c2)::ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PLambda (na1,c1,d1), Lambda(na2,c2,d2) ->
- sorec ((na1,na2,c2)::ctx) (Environ.push_rel (LocalAssum (na2,c2)) env)
+ sorec ((na1,na2,c2)::ctx) (EConstr.push_rel (LocalAssum (na2,c2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PLetIn (na1,c1,Some t1,d1), LetIn(na2,c2,t2,d2) ->
- sorec ((na1,na2,t2)::ctx) (Environ.push_rel (LocalDef (na2,c2,t2)) env)
+ sorec ((na1,na2,t2)::ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env (sorec ctx env subst c1 c2) t1 t2)) d1 d2
| PLetIn (na1,c1,None,d1), LetIn(na2,c2,t2,d2) ->
- sorec ((na1,na2,t2)::ctx) (Environ.push_rel (LocalDef (na2,c2,t2)) env)
+ sorec ((na1,na2,t2)::ctx) (EConstr.push_rel (LocalDef (na2,c2,t2)) env)
(add_binders na1 na2 binding_vars (sorec ctx env subst c1 c2)) d1 d2
| PIf (a1,b1,b1'), Case (ci,_,a2,[|b2;b2'|]) ->
- let ctx_b2,b2 = decompose_lam_n_decls ci.ci_cstr_ndecls.(0) b2 in
- let ctx_b2',b2' = decompose_lam_n_decls ci.ci_cstr_ndecls.(1) b2' in
+ let ctx_b2,b2 = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(0) b2 in
+ let ctx_b2',b2' = decompose_lam_n_decls sigma ci.ci_cstr_ndecls.(1) b2' in
let n = Context.Rel.length ctx_b2 in
let n' = Context.Rel.length ctx_b2' in
- if noccur_between 1 n b2 && noccur_between 1 n' b2' then
+ if Vars.noccur_between sigma 1 n b2 && Vars.noccur_between sigma 1 n' b2' then
let f l (LocalAssum (na,t) | LocalDef (na,_,t)) = (Anonymous,na,t)::l in
let ctx_br = List.fold_left f ctx ctx_b2 in
let ctx_br' = List.fold_left f ctx ctx_b2' in
let b1 = lift_pattern n b1 and b1' = lift_pattern n' b1' in
- sorec ctx_br' (Environ.push_rel_context ctx_b2' env)
- (sorec ctx_br (Environ.push_rel_context ctx_b2 env)
+ sorec ctx_br' (push_rel_context ctx_b2' env)
+ (sorec ctx_br (push_rel_context ctx_b2 env)
(sorec ctx env subst a1 a2) b1 b2) b1' b2'
else
raise PatternMatchingFailure
@@ -310,8 +321,8 @@ let matches_core env sigma convert allow_partial_app allow_bound_rels
let chk_head = sorec ctx env (sorec ctx env subst a1 a2) p1 p2 in
List.fold_left chk_branch chk_head br1
- | PFix c1, Fix _ when eq_constr (mkFix c1) cT -> subst
- | PCoFix c1, CoFix _ when eq_constr (mkCoFix c1) cT -> subst
+ | PFix c1, Fix _ when eq_constr sigma (mkFix (to_fix c1)) cT -> subst
+ | PCoFix c1, CoFix _ when eq_constr sigma (mkCoFix (to_fix c1)) cT -> subst
| _ -> raise PatternMatchingFailure
in
@@ -337,8 +348,9 @@ let mkresult s c n = IStream.Cons ( { m_sub=s; m_ctx=c; } , (IStream.thunk n) )
let isPMeta = function PMeta _ -> true | _ -> false
let matches_head env sigma pat c =
+ let open EConstr in
let head =
- match pat, kind_of_term c with
+ match pat, EConstr.kind sigma c with
| PApp (c1,arg1), App (c2,arg2) ->
if isPMeta c1 then c else
let n1 = Array.length arg1 in
@@ -351,7 +363,7 @@ let matches_head env sigma pat c =
let authorized_occ env sigma partial_app closed pat c mk_ctx =
try
let subst = matches_core_closed env sigma false partial_app pat c in
- if closed && Id.Map.exists (fun _ c -> not (closed0 c)) (snd subst)
+ if closed && Id.Map.exists (fun _ c -> not (closed0 sigma c)) (snd subst)
then (fun next -> next ())
else (fun next -> mkresult subst (mk_ctx (mkMeta special_meta)) next)
with PatternMatchingFailure -> (fun next -> next ())
@@ -360,9 +372,10 @@ let subargs env v = Array.map_to_list (fun c -> (env, c)) v
(* Tries to match a subterm of [c] with [pat] *)
let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
+ let open EConstr in
let rec aux env c mk_ctx next =
let here = authorized_occ env sigma partial_app closed pat c mk_ctx in
- let next () = match kind_of_term c with
+ let next () = match EConstr.kind sigma c with
| Cast (c1,k,c2) ->
let next_mk_ctx = function
| [c1] -> mk_ctx (mkCast (c1, k, c2))
@@ -374,21 +387,21 @@ let sub_match ?(partial_app=false) ?(closed=true) env sigma pat c =
| [c1; c2] -> mk_ctx (mkLambda (x, c1, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalAssum (x,c1)) env in
+ let env' = EConstr.push_rel (LocalAssum (x,c1)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| Prod (x,c1,c2) ->
let next_mk_ctx = function
| [c1; c2] -> mk_ctx (mkProd (x, c1, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalAssum (x,c1)) env in
+ let env' = EConstr.push_rel (LocalAssum (x,c1)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| LetIn (x,c1,t,c2) ->
let next_mk_ctx = function
| [c1; c2] -> mk_ctx (mkLetIn (x, c1, t, c2))
| _ -> assert false
in
- let env' = Environ.push_rel (LocalDef (x,c1,t)) env in
+ let env' = EConstr.push_rel (LocalDef (x,c1,t)) env in
try_aux [(env, c1); (env', c2)] next_mk_ctx next
| App (c1,lc) ->
let topdown = true in
diff --git a/pretyping/constr_matching.mli b/pretyping/constr_matching.mli
index ee6c5141b..4734c90a8 100644
--- a/pretyping/constr_matching.mli
+++ b/pretyping/constr_matching.mli
@@ -10,6 +10,7 @@
open Names
open Term
+open EConstr
open Environ
open Pattern
@@ -63,7 +64,7 @@ val matches_conv : env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
(whose hole is denoted here with [special_meta]) *)
type matching_result =
{ m_sub : bound_ident_map * patvar_map;
- m_ctx : constr }
+ m_ctx : EConstr.t }
(** [match_subterm n pat c] returns the substitution and the context
corresponding to each **closed** subterm of [c] matching [pat]. *)
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 0b5ea86d5..8ba408679 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -6,14 +6,17 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Term
+open Environ
+open EConstr
open Vars
open Inductiveops
-open Environ
open Glob_term
open Glob_ops
open Termops
@@ -188,7 +191,7 @@ let _ = declare_bool_option
(* Auxiliary function for MutCase printing *)
(* [computable] tries to tell if the predicate typing the result is inferable*)
-let computable p k =
+let computable sigma p k =
(* We first remove as many lambda as the arity, then we look
if it remains a lambda for a dependent elimination. This function
works for normal eta-expanded term. For non eta-expanded or
@@ -205,29 +208,29 @@ let computable p k =
sinon on perd la réciprocité de la synthèse (qui, lui,
engendrera un prédicat non dépendant) *)
- let sign,ccl = decompose_lam_assum p in
+ let sign,ccl = decompose_lam_assum sigma p in
Int.equal (Context.Rel.length sign) (k + 1)
&&
- noccur_between 1 (k+1) ccl
+ noccur_between sigma 1 (k+1) ccl
-let lookup_name_as_displayed env t s =
- let rec lookup avoid n c = match kind_of_term c with
+let lookup_name_as_displayed env sigma t s =
+ let rec lookup avoid n c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in RenamingForGoal avoid name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal avoid name c' with
| (Name id,avoid') -> if Id.equal id s then Some n else lookup avoid' (n+1) c'
| (Anonymous,avoid') -> lookup avoid' (n+1) (pop c'))
| Cast (c,_,_) -> lookup avoid n c
| _ -> None
in lookup (ids_of_named_context (named_context env)) 1 t
-let lookup_index_as_renamed env t n =
- let rec lookup n d c = match kind_of_term c with
+let lookup_index_as_renamed env sigma t n =
+ let rec lookup n d c = match EConstr.kind sigma c with
| Prod (name,_,c') ->
- (match compute_displayed_name_in RenamingForGoal [] name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal [] name c' with
(Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -237,7 +240,7 @@ let lookup_index_as_renamed env t n =
else
lookup (n-1) (d+1) c')
| LetIn (name,_,_,c') ->
- (match compute_displayed_name_in RenamingForGoal [] name c' with
+ (match compute_displayed_name_in sigma RenamingForGoal [] name c' with
| (Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) ->
if Int.equal n 0 then
@@ -254,20 +257,20 @@ let lookup_index_as_renamed env t n =
(**********************************************************************)
(* Fragile algorithm to reverse pattern-matching compilation *)
-let update_name na ((_,(e,_)),c) =
+let update_name sigma na ((_,(e,_)),c) =
match na with
- | Name _ when force_wildcard () && noccurn (List.index Name.equal na e) c ->
+ | Name _ when force_wildcard () && noccurn sigma (List.index Name.equal na e) c ->
Anonymous
| _ ->
na
-let rec decomp_branch tags nal b (avoid,env as e) c =
+let rec decomp_branch tags nal b (avoid,env as e) sigma c =
let flag = if b then RenamingForGoal else RenamingForCasesPattern (fst env,c) in
match tags with
| [] -> (List.rev nal,(e,c))
| b::tags ->
let na,c,f,body,t =
- match kind_of_term (strip_outer_cast c), b with
+ match EConstr.kind sigma (strip_outer_cast sigma c), b with
| Lambda (na,t,c),false -> na,c,compute_displayed_let_name_in,None,Some t
| LetIn (na,b,t,c),true ->
na,c,compute_displayed_name_in,Some b,Some t
@@ -277,52 +280,52 @@ let rec decomp_branch tags nal b (avoid,env as e) c =
| _, true ->
Anonymous,lift 1 c,compute_displayed_name_in,None,None
in
- let na',avoid' = f flag avoid na c in
+ let na',avoid' = f sigma flag avoid na c in
decomp_branch tags (na'::nal) b
- (avoid', add_name_opt na' body t env) c
+ (avoid', add_name_opt na' body t env) sigma c
-let rec build_tree na isgoal e ci cl =
- let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in
+let rec build_tree na isgoal e sigma ci cl =
+ let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name sigma na rhs) in
let cnl = ci.ci_pp_info.cstr_tags in
let cna = ci.ci_cstr_nargs in
List.flatten
(List.init (Array.length cl)
- (fun i -> contract_branch isgoal e (cnl.(i),cna.(i),mkpat i,cl.(i))))
+ (fun i -> contract_branch isgoal e sigma (cnl.(i),cna.(i),mkpat i,cl.(i))))
-and align_tree nal isgoal (e,c as rhs) = match nal with
+and align_tree nal isgoal (e,c as rhs) sigma = match nal with
| [] -> [[],rhs]
| na::nal ->
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Case (ci,p,c,cl) when
- eq_constr c (mkRel (List.index Name.equal na (fst (snd e))))
+ eq_constr sigma c (mkRel (List.index Name.equal na (fst (snd e))))
&& not (Int.equal (Array.length cl) 0)
&& (* don't contract if p dependent *)
- computable p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) ->
- let clauses = build_tree na isgoal e ci cl in
+ computable sigma p (List.length ci.ci_pp_info.ind_tags) (* FIXME: can do better *) ->
+ let clauses = build_tree na isgoal e sigma ci cl in
List.flatten
(List.map (fun (pat,rhs) ->
- let lines = align_tree nal isgoal rhs in
+ let lines = align_tree nal isgoal rhs sigma in
List.map (fun (hd,rest) -> pat::hd,rest) lines)
clauses)
| _ ->
- let pat = PatVar(dl,update_name na rhs) in
- let mat = align_tree nal isgoal rhs in
+ let pat = PatVar(dl,update_name sigma na rhs) in
+ let mat = align_tree nal isgoal rhs sigma in
List.map (fun (hd,rest) -> pat::hd,rest) mat
-and contract_branch isgoal e (cdn,can,mkpat,b) =
- let nal,rhs = decomp_branch cdn [] isgoal e b in
- let mat = align_tree nal isgoal rhs in
+and contract_branch isgoal e sigma (cdn,can,mkpat,b) =
+ let nal,rhs = decomp_branch cdn [] isgoal e sigma b in
+ let mat = align_tree nal isgoal rhs sigma in
List.map (fun (hd,rhs) -> (mkpat rhs hd,rhs)) mat
(**********************************************************************)
(* Transform internal representation of pattern-matching into list of *)
(* clauses *)
-let is_nondep_branch c l =
+let is_nondep_branch sigma c l =
try
(* FIXME: do better using tags from l *)
- let sign,ccl = decompose_lam_n_decls (List.length l) c in
- noccur_between 1 (Context.Rel.length sign) ccl
+ let sign,ccl = decompose_lam_n_decls sigma (List.length l) c in
+ noccur_between sigma 1 (Context.Rel.length sign) ccl
with e when CErrors.noncritical e -> (* Not eta-expanded or not reduced *)
false
@@ -420,7 +423,7 @@ let detype_sort sigma = function
| Type u ->
GType
(if !print_universes
- then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Evd.pr_evd_level sigma) u)]
+ then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)]
else [])
type binder_kind = BProd | BLambda | BLetIn
@@ -432,14 +435,15 @@ let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index
let set_detype_anonymous f = detype_anonymous := f
let detype_level sigma l =
- GType (Some (dl, Pp.string_of_ppcmds (Evd.pr_evd_level sigma l)))
+ GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l)))
let detype_instance sigma l =
+ let l = EInstance.kind sigma l in
if Univ.Instance.is_empty l then None
else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l)))
let rec detype flags avoid env sigma t =
- match kind_of_term (collapse_appl t) with
+ match EConstr.kind sigma (collapse_appl sigma t) with
| Rel n ->
(try match lookup_name_of_rel n (fst env) with
| Name id -> GVar (dl, id)
@@ -454,7 +458,7 @@ let rec detype flags avoid env sigma t =
| Var id ->
(try let _ = Global.lookup_named id in GRef (dl, VarRef id, None)
with Not_found -> GVar (dl, id))
- | Sort s -> GSort (dl,detype_sort sigma s)
+ | Sort s -> GSort (dl,detype_sort sigma (ESorts.kind sigma s))
| Cast (c1,REVERTcast,c2) when not !Flags.raw_print ->
detype flags avoid env sigma c1
| Cast (c1,k,c2) ->
@@ -504,7 +508,9 @@ let rec detype flags avoid env sigma t =
let ty = Retyping.get_type_of (snd env) sigma c in
let ((ind,u), args) = Inductiveops.find_mrectype (snd env) sigma ty in
let body' = strip_lam_assum body in
- let body' = subst_instance_constr u body' in
+ let u = EInstance.kind sigma u in
+ let body' = CVars.subst_instance_constr u body' in
+ let body' = EConstr.of_constr body' in
substl (c :: List.rev args) body'
with Retyping.RetypeError _ | Not_found ->
anomaly (str"Cannot detype an unfolded primitive projection.")
@@ -523,18 +529,18 @@ let rec detype flags avoid env sigma t =
| LocalDef _ -> true
| LocalAssum (id,_) ->
try let n = List.index Name.equal (Name id) (fst env) in
- isRelN n c
- with Not_found -> isVarId id c
+ isRelN sigma n c
+ with Not_found -> isVarId sigma id c
in
let id,l =
try
let id = match Evd.evar_ident evk sigma with
- | None -> Evd.pr_evar_suggested_name evk sigma
+ | None -> Termops.pr_evar_suggested_name evk sigma
| Some id -> id
in
let l = Evd.evar_instance_array bound_to_itself_or_letin (Evd.find sigma evk) cl in
- let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match kind_of_term c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in
- let l = Evd.evar_instance_array (fun d c -> not !print_evar_arguments && (bound_to_itself_or_letin d c && not (isRel c && Int.Set.mem (destRel c) rels || isVar c && (Id.Set.mem (destVar c) fvs)))) (Evd.find sigma evk) cl in
+ let fvs,rels = List.fold_left (fun (fvs,rels) (_,c) -> match EConstr.kind sigma c with Rel n -> (fvs,Int.Set.add n rels) | Var id -> (Id.Set.add id fvs,rels) | _ -> (fvs,rels)) (Id.Set.empty,Int.Set.empty) l in
+ 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)),
@@ -547,10 +553,10 @@ let rec detype flags avoid env sigma t =
| Construct (cstr_sp,u) ->
GRef (dl, ConstructRef cstr_sp, detype_instance sigma u)
| Case (ci,p,c,bl) ->
- let comp = computable p (List.length (ci.ci_pp_info.ind_tags)) in
+ let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in
detype_case comp (detype flags avoid env sigma)
(detype_eqns flags avoid env sigma ci comp)
- is_nondep_branch avoid
+ (is_nondep_branch sigma) avoid
(ci.ci_ind,ci.ci_pp_info.style,
ci.ci_pp_info.cstr_tags,ci.ci_pp_info.ind_tags)
(Some p) c bl
@@ -590,7 +596,7 @@ and detype_cofix flags avoid env sigma n (names,tys,bodies) =
Array.map (fun (_,bd,_) -> bd) v)
and share_names flags n l avoid env sigma c t =
- match kind_of_term c, kind_of_term t with
+ match EConstr.kind sigma c, EConstr.kind sigma t with
(* factorize even when not necessary to have better presentation *)
| Lambda (na,t,c), Prod (na',t',c') ->
let na = match (na,na') with
@@ -628,7 +634,7 @@ and share_names flags n l avoid env sigma c t =
and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
try
if !Flags.raw_print || not (reverse_matching ()) then raise Exit;
- let mat = build_tree Anonymous (snd flags) (avoid,env) ci bl in
+ let mat = build_tree Anonymous (snd flags) (avoid,env) sigma ci bl in
List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype flags avoid env sigma c))
mat
with e when CErrors.noncritical e ->
@@ -637,15 +643,15 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch =
let make_pat x avoid env b body ty ids =
- if force_wildcard () && noccurn 1 b then
+ if force_wildcard () && noccurn sigma 1 b then
PatVar (dl,Anonymous),avoid,(add_name Anonymous body ty env),ids
else
let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in
- let na,avoid' = compute_displayed_name_in flag avoid x b in
+ let na,avoid' = compute_displayed_name_in sigma flag avoid x b in
PatVar (dl,na),avoid',(add_name na body ty env),add_vname ids na
in
let rec buildrec ids patlist avoid env l b =
- match kind_of_term b, l with
+ match EConstr.kind sigma b, l with
| _, [] ->
(dl, Id.Set.elements ids,
[PatCstr(dl, constr, List.rev patlist,Anonymous)],
@@ -680,8 +686,8 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c =
let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in
let na',avoid' = match bk with
- | BLetIn -> compute_displayed_let_name_in flag avoid na c
- | _ -> compute_displayed_name_in flag avoid na c in
+ | BLetIn -> compute_displayed_let_name_in sigma flag avoid na c
+ | _ -> compute_displayed_name_in sigma flag avoid na c in
let r = detype flags avoid' (add_name na' body ty env) sigma c in
match bk with
| BProd -> GProd (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r)
@@ -694,7 +700,7 @@ and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c =
GLetIn (dl, na', c, t, r)
let detype_rel_context ?(lax=false) where avoid env sigma sign =
- let where = Option.map (fun c -> it_mkLambda_or_LetIn c sign) where in
+ let where = Option.map (fun c -> EConstr.it_mkLambda_or_LetIn c sign) where in
let rec aux avoid env = function
| [] -> []
| decl::rest ->
@@ -706,10 +712,10 @@ let detype_rel_context ?(lax=false) where avoid env sigma sign =
| None -> na,avoid
| Some c ->
if is_local_def decl then
- compute_displayed_let_name_in
+ compute_displayed_let_name_in sigma
(RenamingElsewhereFor (fst env,c)) avoid na c
else
- compute_displayed_name_in
+ compute_displayed_name_in sigma
(RenamingElsewhereFor (fst env,c)) avoid na c in
let b = match decl with
| LocalAssum _ -> None
@@ -726,6 +732,7 @@ let detype ?(lax=false) isgoal avoid env sigma t =
detype (lax,isgoal) avoid (names_of_rel_context env, env) sigma t
let detype_closed_glob ?lax isgoal avoid env sigma t =
+ let open Context.Rel.Declaration in
let convert_id cl id =
try Id.Map.find id cl.idents
with Not_found -> id
@@ -746,8 +753,8 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
(* spiwack: I'm not sure it is the right thing to do,
but I'm computing the detyping environment like
[Printer.pr_constr_under_binders_env] does. *)
- let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) b in
- let env = Termops.push_rels_assum assums env in
+ let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in
+ let env = push_rel_context assums env in
detype ?lax isgoal avoid env sigma c
(* if [id] is bound to a [closed_glob_constr]. *)
with Not_found -> try
@@ -802,7 +809,7 @@ let rec subst_glob_constr subst raw =
| GRef (loc,ref,u) ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
- detype false [] (Global.env()) Evd.empty t
+ detype false [] (Global.env()) Evd.empty (EConstr.of_constr t)
| GVar _ -> raw
| GEvar _ -> raw
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index c51cb0f44..4c6f9129f 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -9,6 +9,7 @@
open Names
open Term
open Environ
+open EConstr
open Glob_term
open Termops
open Mod_subst
@@ -45,13 +46,13 @@ val detype_case :
val detype_sort : evar_map -> sorts -> glob_sort
val detype_rel_context : ?lax:bool -> constr option -> Id.t list -> (names_context * env) ->
- evar_map -> Context.Rel.t -> glob_decl list
+ evar_map -> rel_context -> glob_decl list
val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> closed_glob_constr -> glob_constr
(** look for the index of a named var or a nondep var as it is renamed *)
-val lookup_name_as_displayed : env -> constr -> Id.t -> int option
-val lookup_index_as_renamed : env -> constr -> int -> int option
+val lookup_name_as_displayed : env -> evar_map -> constr -> Id.t -> int option
+val lookup_index_as_renamed : env -> evar_map -> constr -> int -> int option
val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit
val force_wildcard : unit -> bool
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index d18b437a3..4bb66b8e9 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -10,12 +10,13 @@ open CErrors
open Util
open Names
open Term
+open Termops
+open Environ
+open EConstr
open Vars
open CClosure
open Reduction
open Reductionops
-open Termops
-open Environ
open Recordops
open Evarutil
open Evardefine
@@ -30,7 +31,7 @@ module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
type unify_fun = transparent_state ->
- env -> evar_map -> conv_pb -> constr -> constr -> Evarsolve.unification_result
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> Evarsolve.unification_result
let debug_unification = ref (false)
let _ = Goptions.declare_bool_option {
@@ -49,10 +50,10 @@ let unfold_projection env evd ts p c =
else None
let eval_flexible_term ts env evd c =
- match kind_of_term c with
- | Const (c,u as cu) ->
+ match EConstr.kind evd c with
+ | Const (c, u) ->
if is_transparent_constant ts c
- then constant_opt_value_in env cu
+ then Option.map EConstr.of_constr (constant_opt_value_in env (c, EInstance.kind evd u))
else None
| Rel n ->
(try match lookup_rel n env with
@@ -74,11 +75,11 @@ let eval_flexible_term ts env evd c =
type flex_kind_of_term =
| Rigid
- | MaybeFlexible of Constr.t (* reducible but not necessarily reduced *)
- | Flexible of existential
+ | MaybeFlexible of EConstr.t (* reducible but not necessarily reduced *)
+ | Flexible of EConstr.existential
let flex_kind_of_term ts env evd c sk =
- match kind_of_term c with
+ match EConstr.kind evd c with
| LetIn _ | Rel _ | Const _ | Var _ | Proj _ ->
Option.cata (fun x -> MaybeFlexible x) Rigid (eval_flexible_term ts env evd c)
| Lambda _ when not (Option.is_empty (Stack.decomp sk)) -> MaybeFlexible c
@@ -88,10 +89,13 @@ let flex_kind_of_term ts env evd c sk =
| Fix _ -> Rigid (* happens when the fixpoint is partially applied *)
| Cast _ | App _ | Case _ -> assert false
+let add_conv_pb (pb, env, x, y) sigma =
+ Evd.add_conv_pb (pb, env, EConstr.Unsafe.to_constr x, EConstr.Unsafe.to_constr y) sigma
+
let apprec_nohdbeta ts env evd c =
let (t,sk as appr) = Reductionops.whd_nored_state evd (c, []) in
if Stack.not_purely_applicative sk
- then Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state
+ then Stack.zip evd (fst (whd_betaiota_deltazeta_for_iota_state
ts env evd Cst_stack.empty appr))
else c
@@ -101,7 +105,7 @@ let position_problem l2r = function
let occur_rigidly (evk,_ as ev) evd t =
let rec aux t =
- match kind_of_term (whd_evar evd t) with
+ match EConstr.kind evd t with
| App (f, c) -> if aux f then Array.exists aux c else false
| Construct _ | Ind _ | Sort _ | Meta _ | Fix _ | CoFix _ -> true
| Proj (p, c) -> not (aux c)
@@ -111,7 +115,7 @@ let occur_rigidly (evk,_ as ev) evd t =
| Const _ -> false
| Prod (_, b, t) -> ignore(aux b || aux t); true
| Rel _ | Var _ -> false
- | Case (_,_,c,_) -> if eq_constr (mkEvar ev) c then raise Occur else false
+ | Case (_,_,c,_) -> if eq_constr evd (mkEvar ev) c then raise Occur else false
in try ignore(aux t); false with Occur -> true
(* [check_conv_record env sigma (t1,stack1) (t2,stack2)] tries to decompose
@@ -135,20 +139,22 @@ let occur_rigidly (evk,_ as ev) evd t =
projection would have been reduced) *)
let check_conv_record env sigma (t1,sk1) (t2,sk2) =
- let (proji, u), arg = Universes.global_app_of_constr t1 in
+ let (proji, u), arg = Termops.global_app_of_constr sigma t1 in
let canon_s,sk2_effective =
try
- match kind_of_term t2 with
+ match EConstr.kind sigma t2 with
Prod (_,a,b) -> (* assert (l2=[]); *)
- let _, a, b = destProd (Evarutil.nf_evar sigma t2) in
- if dependent (mkRel 1) b then raise Not_found
- else lookup_canonical_conversion (proji, Prod_cs),
+ let _, a, b = destProd sigma t2 in
+ if noccurn sigma 1 b then
+ lookup_canonical_conversion (proji, Prod_cs),
(Stack.append_app [|a;pop b|] Stack.empty)
+ else raise Not_found
| Sort s ->
+ let s = ESorts.kind sigma s in
lookup_canonical_conversion
(proji, Sort_cs (family_of_sort s)),[]
| _ ->
- let c2 = global_of_constr t2 in
+ let (c2, _) = Termops.global_of_constr sigma t2 in
lookup_canonical_conversion (proji, Const_cs c2),sk2
with Not_found ->
let (c, cs) = lookup_canonical_conversion (proji,Default_cs) in
@@ -156,12 +162,14 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
in
let t', { o_DEF = c; o_CTX = ctx; o_INJ=n; o_TABS = bs;
o_TPARAMS = params; o_NPARAMS = nparams; o_TCOMPS = us } = canon_s in
+ let us = List.map EConstr.of_constr us in
+ let params = List.map EConstr.of_constr params in
let params1, c1, extra_args1 =
match arg with
| Some c -> (* A primitive projection applied to c *)
let ty = Retyping.get_type_of ~lax:true env sigma c in
let (i,u), ind_args =
- try Inductiveops.find_mrectype env sigma ty
+ try Inductiveops.find_mrectype env sigma ty
with _ -> raise Not_found
in Stack.append_app_list ind_args Stack.empty, c, sk1
| None ->
@@ -175,13 +183,15 @@ let check_conv_record env sigma (t1,sk1) (t2,sk2) =
| None -> raise Not_found
| Some (l',el,s') -> (l'@Stack.append_app [|el|] Stack.empty,s') in
let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in
+ let c = EConstr.of_constr c in
let c' = subst_univs_level_constr subst c in
+ let t' = EConstr.of_constr t' in
let t' = subst_univs_level_constr subst t' in
- let bs' = List.map (subst_univs_level_constr subst) bs in
- let h, _ = decompose_app_vect t' in
+ let bs' = List.map (EConstr.of_constr %> subst_univs_level_constr subst) bs in
+ let h, _ = decompose_app_vect sigma t' in
ctx',(h, t2),c',bs',(Stack.append_app_list params Stack.empty,params1),
(Stack.append_app_list us Stack.empty,us2),(extra_args1,extra_args2),c1,
- (n,Stack.zip(t2,sk2))
+ (n, Stack.zip sigma (t2,sk2))
(* Precondition: one of the terms of the pb is an uninstantiated evar,
* possibly applied to arguments. *)
@@ -234,6 +244,11 @@ let rec ise_app_stack2 env f evd sk1 sk2 =
end
| _, _ -> (sk1,sk2), Success evd
+let push_rec_types pfix env =
+ let (i, c, t) = pfix in
+ let inj c = EConstr.Unsafe.to_constr c in
+ push_rec_types (i, Array.map inj c, Array.map inj t) env
+
(* This function tries to unify 2 stacks element by element. It works
from the end to the beginning. If it unifies a non empty suffix of
stacks but not the entire stacks, the first part of the answer is
@@ -323,7 +338,7 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
let e =
try
let evd, b = infer_conv ~catch_incon:false ~pb:pbty ~ts:(fst ts)
- env evd term1 term2
+ env evd term1 term2
in
if b then Success evd
else UnifFailure (evd, ConversionFailed (env,term1,term2))
@@ -346,16 +361,16 @@ let rec evar_conv_x ts env evd pbty term1 term2 =
(whd_nored_state evd (term1,Stack.empty), Cst_stack.empty)
(whd_nored_state evd (term2,Stack.empty), Cst_stack.empty)
in
- begin match kind_of_term term1, kind_of_term term2 with
+ begin match EConstr.kind evd term1, EConstr.kind evd term2 with
| Evar ev, _ when Evd.is_undefined evd (fst ev) ->
(match solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem true pbty,ev,term2) with
+ (position_problem true pbty,ev, term2) with
| UnifFailure (_,OccurCheck _) ->
(* Eta-expansion might apply *) default ()
| x -> x)
| _, Evar ev when Evd.is_undefined evd (fst ev) ->
(match solve_simple_eqn (evar_conv_x ts) env evd
- (position_problem false pbty,ev,term1) with
+ (position_problem false pbty,ev, term1) with
| UnifFailure (_, OccurCheck _) ->
(* Eta-expansion might apply *) default ()
| x -> x)
@@ -371,8 +386,8 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
match is_unification_pattern_evar env evd ev lF tM with
| None -> fallback ()
| Some l1' -> (* Miller-Pfenning's patterns unification *)
- let t2 = nf_evar evd tM in
- let t2 = solve_pattern_eqn env l1' t2 in
+ let t2 = tM in
+ let t2 = solve_pattern_eqn env evd l1' t2 in
solve_simple_eqn (evar_conv_x ts) env evd
(position_problem on_left pbty,ev,t2)
in
@@ -381,27 +396,27 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let not_only_app = Stack.not_purely_applicative skO in
match switch (ise_stack2 not_only_app env evd (evar_conv_x ts)) skF skO with
|Some (l,r), Success i' when on_left && (not_only_app || List.is_empty l) ->
- switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r))
+ switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r))
|Some (r,l), Success i' when not on_left && (not_only_app || List.is_empty l) ->
- switch (evar_conv_x ts env i' pbty) (Stack.zip(termF,l)) (Stack.zip(termO,r))
+ switch (evar_conv_x ts env i' pbty) (Stack.zip evd (termF,l)) (Stack.zip evd (termO,r))
|None, Success i' -> switch (evar_conv_x ts env i' pbty) termF termO
|_, (UnifFailure _ as x) -> x
|Some _, _ -> UnifFailure (evd,NotSameArgSize) in
let eta env evd onleft sk term sk' term' =
assert (match sk with [] -> true | _ -> false);
- let (na,c1,c'1) = destLambda term in
+ let (na,c1,c'1) = destLambda evd term in
let c = nf_evar evd c1 in
let env' = push_rel (RelDecl.LocalAssum (na,c)) env in
let out1 = whd_betaiota_deltazeta_for_iota_state
(fst ts) env' evd Cst_stack.empty (c'1, Stack.empty) in
let out2 = whd_nored_state evd
- (Stack.zip (term', sk' @ [Stack.Shift 1]), Stack.append_app [|mkRel 1|] Stack.empty),
+ (Stack.zip evd (term', sk' @ [Stack.Shift 1]), Stack.append_app [|EConstr.mkRel 1|] Stack.empty),
Cst_stack.empty in
if onleft then evar_eqappr_x ts env' evd CONV out1 out2
else evar_eqappr_x ts env' evd CONV out2 out1
in
let rigids env evd sk term sk' term' =
- let univs = Universes.eq_constr_universes term term' in
+ let univs = EConstr.eq_constr_universes evd term term' in
match univs with
| Some univs ->
ise_and evd [(fun i ->
@@ -419,10 +434,10 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
match Stack.list_of_app_stack skF with
| None -> quick_fail evd
| Some lF ->
- let tM = Stack.zip apprM in
+ let tM = Stack.zip evd apprM in
miller_pfenning on_left
(fun () -> if not_only_app then (* Postpone the use of an heuristic *)
- switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip apprF) tM
+ switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i)) (Stack.zip evd apprF) tM
else quick_fail i)
ev lF tM i
and consume (termF,skF as apprF) (termM,skM as apprM) i =
@@ -436,7 +451,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
in
let default i = ise_try i [f1; consume apprF apprM; delta]
in
- match kind_of_term termM with
+ match EConstr.kind evd termM with
| Proj (p, c) when not (Stack.is_empty skF) ->
(* Might be ?X args = p.c args', and we have to eta-expand the
primitive projection if |args| >= |args'|+1. *)
@@ -466,7 +481,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let flex_rigid on_left ev (termF, skF as apprF) (termR, skR as apprR) =
let switch f a b = if on_left then f a b else f b a in
let eta evd =
- match kind_of_term termR with
+ match EConstr.kind evd termR with
| Lambda _ when (* if ever problem is ill-typed: *) List.is_empty skR ->
eta env evd false skR termR skF termF
| Construct u -> eta_constructor ts env evd skR u skF termF
@@ -476,7 +491,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
| None ->
ise_try evd [consume_stack on_left apprF apprR; eta]
| Some lF ->
- let tR = Stack.zip apprR in
+ let tR = Stack.zip evd apprR in
miller_pfenning on_left
(fun () ->
ise_try evd
@@ -484,12 +499,12 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(fun i ->
if not (occur_rigidly ev i tR) then
let i,tF =
- if isRel tR || isVar tR then
+ if isRel i tR || isVar i tR then
(* Optimization so as to generate candidates *)
let i,ev = evar_absorb_arguments env i ev lF in
i,mkEvar ev
else
- i,Stack.zip apprF in
+ i,Stack.zip evd apprF in
switch (fun x y -> Success (add_conv_pb (pbty,env,x,y) i))
tF tR
else
@@ -512,9 +527,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We do have sk1[] = sk2[]: we now unify ?ev1 and ?ev2 *)
(* Note that ?ev1 and ?ev2, may have been instantiated in the meantime *)
let ev1' = whd_evar i' (mkEvar ev1) in
- if isEvar ev1' then
+ if isEvar i' ev1' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem true pbty,destEvar ev1',term2)
+ (position_problem true pbty,destEvar i' ev1', term2)
else
evar_eqappr_x ts env evd pbty
((ev1', sk1), csts1) ((term2, sk2), csts2)
@@ -522,9 +537,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We have sk1'[] = sk2[] for some sk1' s.t. sk1[]=sk1'[r[]] *)
(* we now unify r[?ev1] and ?ev2 *)
let ev2' = whd_evar i' (mkEvar ev2) in
- if isEvar ev2' then
+ if isEvar i' ev2' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem false pbty,destEvar ev2',Stack.zip(term1,r))
+ (position_problem false pbty,destEvar i' ev2',Stack.zip evd (term1,r))
else
evar_eqappr_x ts env evd pbty
((ev2', sk1), csts1) ((term2, sk2), csts2)
@@ -533,9 +548,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* We have sk1[] = sk2'[] for some sk2' s.t. sk2[]=sk2'[r[]] *)
(* we now unify ?ev1 and r[?ev2] *)
let ev1' = whd_evar i' (mkEvar ev1) in
- if isEvar ev1' then
+ if isEvar i' ev1' then
solve_simple_eqn (evar_conv_x ts) env i'
- (position_problem true pbty,destEvar ev1',Stack.zip(term2,r))
+ (position_problem true pbty,destEvar i' ev1',Stack.zip evd (term2,r))
else evar_eqappr_x ts env evd pbty
((ev1', sk1), csts1) ((term2, sk2), csts2)
| None, (UnifFailure _ as x) ->
@@ -588,7 +603,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
flex_maybeflex false ev2 (appr2,csts2) (appr1,csts1) v1
| MaybeFlexible v1, MaybeFlexible v2 -> begin
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| LetIn (na1,b1,t1,c'1), LetIn (na2,b2,t2,c'2) ->
let f1 i = (* FO *)
ise_and i
@@ -625,7 +640,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
(* Catch the p.c ~= p c' cases *)
| Proj (p,c), Const (p',u) when eq_constant (Projection.constant p) p' ->
let res =
- try Some (destApp (Retyping.expand_projection env evd p c []))
+ try Some (destApp evd (Retyping.expand_projection env evd p c []))
with Retyping.RetypeError _ -> None
in
(match res with
@@ -636,7 +651,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
| Const (p,u), Proj (p',c') when eq_constant p (Projection.constant p') ->
let res =
- try Some (destApp (Retyping.expand_projection env evd p' c' []))
+ try Some (destApp evd (Retyping.expand_projection env evd p' c' []))
with Retyping.RetypeError _ -> None
in
(match res with
@@ -651,7 +666,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
allow this identification (first-order unification of universes). Otherwise
fallback to unfolding.
*)
- let univs = Universes.eq_constr_universes term1 term2 in
+ let univs = EConstr.eq_constr_universes evd term1 term2 in
match univs with
| Some univs ->
ise_and i [(fun i ->
@@ -673,7 +688,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
if the first argument is a beta-redex (expand a constant
only if necessary) or the second argument is potentially
usable as a canonical projection or canonical value *)
- let rec is_unnamed (hd, args) = match kind_of_term hd with
+ let rec is_unnamed (hd, args) = match EConstr.kind i hd with
| (Var _|Construct _|Ind _|Const _|Prod _|Sort _) ->
Stack.not_purely_applicative args
| (CoFix _|Meta _|Rel _)-> true
@@ -694,7 +709,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
let rhs_is_already_stuck =
rhs_is_already_stuck || rhs_is_stuck_and_unnamed () in
- if (isLambda term1 || rhs_is_already_stuck)
+ if (EConstr.isLambda i term1 || rhs_is_already_stuck)
&& (not (Stack.not_purely_applicative sk1)) then
evar_eqappr_x ~rhs_is_already_stuck ts env i pbty
(whd_betaiota_deltazeta_for_iota_state
@@ -708,9 +723,9 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
ise_try evd [f1; f2; f3]
end
- | Rigid, Rigid when isLambda term1 && isLambda term2 ->
- let (na1,c1,c'1) = destLambda term1 in
- let (na2,c2,c'2) = destLambda term2 in
+ | Rigid, Rigid when EConstr.isLambda evd term1 && EConstr.isLambda evd term2 ->
+ let (na1,c1,c'1) = EConstr.destLambda evd term1 in
+ let (na2,c2,c'2) = EConstr.destLambda evd term2 in
assert app_empty;
ise_and evd
[(fun i -> evar_conv_x ts env i CONV c1 c2);
@@ -750,17 +765,19 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
ise_try evd [f3; f4]
(* Eta-expansion *)
- | Rigid, _ when isLambda term1 && (* if ever ill-typed: *) List.is_empty sk1 ->
+ | Rigid, _ when isLambda evd term1 && (* if ever ill-typed: *) List.is_empty sk1 ->
eta env evd true sk1 term1 sk2 term2
- | _, Rigid when isLambda term2 && (* if ever ill-typed: *) List.is_empty sk2 ->
+ | _, Rigid when isLambda evd term2 && (* if ever ill-typed: *) List.is_empty sk2 ->
eta env evd false sk2 term2 sk1 term1
| Rigid, Rigid -> begin
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| Sort s1, Sort s2 when app_empty ->
(try
+ let s1 = ESorts.kind evd s1 in
+ let s2 = ESorts.kind evd s2 in
let evd' =
if pbty == CONV
then Evd.set_eq_sort env evd s1 s2
@@ -823,7 +840,7 @@ and evar_eqappr_x ?(rhs_is_already_stuck = false) ts env evd pbty
begin match ise_stack2 true env evd (evar_conv_x ts) sk1 sk2 with
|_, (UnifFailure _ as x) -> x
|None, Success i' -> evar_conv_x ts env i' CONV term1 term2
- |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip (term1,sk1')) (Stack.zip (term2,sk2'))
+ |Some (sk1',sk2'), Success i' -> evar_conv_x ts env i' CONV (Stack.zip i' (term1,sk1')) (Stack.zip i' (term2,sk2'))
end
| (Ind _ | Sort _ | Prod _ | CoFix _ | Fix _ | Rel _ | Var _ | Const _), _ ->
@@ -892,7 +909,7 @@ and conv_record trs env evd (ctx,(h,h2),c,bs,(params,params1),(us,us2),(sk1,sk2)
(fun i -> exact_ise_stack2 env i (evar_conv_x trs) sk1 sk2);
test;
(fun i -> evar_conv_x trs env i CONV h2
- (fst (decompose_app_vect (substl ks h))))]
+ (fst (decompose_app_vect i (substl ks h))))]
else UnifFailure(evd,(*dummy*)NotSameHead)
and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 =
@@ -903,8 +920,8 @@ and eta_constructor ts env evd sk1 ((ind, i), u) sk2 term2 =
(try
let l1' = Stack.tail pars sk1 in
let l2' =
- let term = Stack.zip (term2,sk2) in
- List.map (fun p -> mkProj (Projection.make p false, term)) (Array.to_list projs)
+ let term = Stack.zip evd (term2,sk2) in
+ List.map (fun p -> EConstr.mkProj (Projection.make p false, term)) (Array.to_list projs)
in
exact_ise_stack2 env evd (evar_conv_x (fst ts, false)) l1'
(Stack.append_app_list l2' Stack.empty)
@@ -948,50 +965,56 @@ let first_order_unification ts env evd (ev1,l1) (term2,l2) =
let choose_less_dependent_instance evk evd term args =
let evi = Evd.find_undefined evd evk in
let subst = make_pure_subst evi args in
- let subst' = List.filter (fun (id,c) -> Term.eq_constr c term) subst in
+ let subst' = List.filter (fun (id,c) -> EConstr.eq_constr evd c term) subst in
match subst' with
| [] -> None
- | (id, _) :: _ -> Some (Evd.define evk (mkVar id) evd)
+ | (id, _) :: _ -> Some (Evd.define evk (Constr.mkVar id) evd)
let apply_on_subterm env evdref f c t =
let rec applyrec (env,(k,c) as acc) t =
(* By using eq_constr, we make an approximation, for instance, we *)
(* could also be interested in finding a term u convertible to t *)
(* such that c occurs in u *)
- if e_eq_constr_univs evdref c t then f k
+ let eq_constr c1 c2 = match EConstr.eq_constr_universes !evdref c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !evdref cstr); true
+ with UniversesDiffer -> false
+ in
+ if eq_constr c t then f k
else
- match kind_of_term t with
- | Evar (evk,args) when Evd.is_undefined !evdref evk ->
+ match EConstr.kind !evdref t with
+ | Evar (evk,args) ->
let ctx = evar_filtered_context (Evd.find_undefined !evdref evk) in
let g decl a = if is_local_assum decl then applyrec acc a else a in
mkEvar (evk, Array.of_list (List.map2 g ctx (Array.to_list args)))
| _ ->
- map_constr_with_binders_left_to_right
+ map_constr_with_binders_left_to_right !evdref
(fun d (env,(k,c)) -> (push_rel d env, (k+1,lift 1 c)))
applyrec acc t
in
applyrec (env,(0,c)) t
-let filter_possible_projections c ty ctxt args =
+let filter_possible_projections evd c ty ctxt args =
(* Since args in the types will be replaced by holes, we count the
fv of args to have a well-typed filter; don't know how necessary
- it is however to have a well-typed filter here *)
- let fv1 = free_rels (mkApp (c,args)) (* Hack: locally untyped *) in
- let fv2 = collect_vars (mkApp (c,args)) in
+ it is however to have a well-typed filter here *)
+ let fv1 = free_rels evd (mkApp (c,args)) (* Hack: locally untyped *) in
+ let fv2 = collect_vars evd (mkApp (c,args)) in
let len = Array.length args in
- let tyvars = collect_vars ty in
+ let tyvars = collect_vars evd ty in
List.map_i (fun i decl ->
let () = assert (i < len) in
let a = Array.unsafe_get args i in
(match decl with
| NamedDecl.LocalAssum _ -> false
- | NamedDecl.LocalDef (_,c,_) -> not (isRel c || isVar c)) ||
+ | NamedDecl.LocalDef (_,c,_) -> not (isRel evd (EConstr.of_constr c) || isVar evd (EConstr.of_constr c))) ||
a == c ||
(* Here we make an approximation, for instance, we could also be *)
(* interested in finding a term u convertible to c such that a occurs *)
(* in u *)
- isRel a && Int.Set.mem (destRel a) fv1 ||
- isVar a && Id.Set.mem (destVar a) fv2 ||
+ isRel evd a && Int.Set.mem (destRel evd a) fv1 ||
+ isVar evd a && Id.Set.mem (destVar evd a) fv2 ||
Id.Set.mem (NamedDecl.get_id decl) tyvars)
0 ctxt
@@ -1026,7 +1049,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
let instance = List.map mkVar (List.map NamedDecl.get_id ctxt) in
let rec make_subst = function
- | decl'::ctxt', c::l, occs::occsl when isVarId (NamedDecl.get_id decl') c ->
+ | decl'::ctxt', c::l, occs::occsl when isVarId evd (NamedDecl.get_id decl') c ->
begin match occs with
| Some _ ->
error "Cannot force abstraction on identity instance."
@@ -1035,10 +1058,10 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
end
| decl'::ctxt', c::l, occs::occsl ->
let id = NamedDecl.get_id decl' in
- let t = NamedDecl.get_type decl' in
+ let t = EConstr.of_constr (NamedDecl.get_type decl') in
let evs = ref [] in
let ty = Retyping.get_type_of env_rhs evd c in
- let filter' = filter_possible_projections c ty ctxt args in
+ let filter' = filter_possible_projections evd c ty ctxt args in
(id,t,c,ty,evs,Filter.make filter',occs) :: make_subst (ctxt',l,occsl)
| _, _, [] -> []
| _ -> anomaly (Pp.str "Signature or instance are shorter than the occurrences list") in
@@ -1056,7 +1079,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
let Sigma (ev, evd, _) = new_evar_instance sign evd evty ~filter instance in
let evd = Sigma.to_evar_map evd in
evdref := evd;
- evsref := (fst (destEvar ev),evty)::!evsref;
+ evsref := (fst (destEvar !evdref ev),evty)::!evsref;
ev in
set_holes evdref (apply_on_subterm env_rhs evdref set_var c rhs) subst
| [] -> rhs in
@@ -1084,7 +1107,7 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
(* We force abstraction over this unconstrained occurrence *)
(* and we use typing to propagate this instantiation *)
(* This is an arbitrary choice *)
- let evd = Evd.define evk (mkVar id) evd in
+ let evd = Evd.define evk (Constr.mkVar id) evd in
match evar_conv_x ts env_evar evd CUMUL idty evty with
| UnifFailure _ -> error "Cannot find an instance"
| Success evd ->
@@ -1102,15 +1125,18 @@ let second_order_matching ts env_rhs evd (evk,args) argoccs rhs =
force_instantiation evd !evsref
| [] ->
let evd =
- try Evarsolve.check_evar_instance evd evk rhs
+ try Evarsolve.check_evar_instance evd evk rhs
(evar_conv_x full_transparent_state)
with IllTypedInstance _ -> raise (TypingFailed evd)
in
- Evd.define evk rhs evd
+ Evd.define evk (EConstr.Unsafe.to_constr rhs) evd
in
abstract_free_holes evd subst, true
with TypingFailed evd -> evd, false
+let to_pb (pb, env, t1, t2) =
+ (pb, env, EConstr.Unsafe.to_constr t1, EConstr.Unsafe.to_constr t2)
+
let second_order_matching_with_args ts env evd pbty ev l t =
(*
let evd,ev = evar_absorb_arguments env evd ev l in
@@ -1120,22 +1146,22 @@ let second_order_matching_with_args ts env evd pbty ev l t =
else UnifFailure (evd, ConversionFailed (env,mkApp(mkEvar ev,l),t))
if b then Success evd else
*)
- let pb = (pbty,env,mkApp(mkEvar ev,l),t) in
+ let pb = to_pb (pbty,env,mkApp(mkEvar ev,l),t) in
UnifFailure (evd, CannotSolveConstraint (pb,ProblemBeyondCapabilities))
let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
let t1 = apprec_nohdbeta ts env evd (whd_head_evar evd t1) in
let t2 = apprec_nohdbeta ts env evd (whd_head_evar evd t2) in
- let (term1,l1 as appr1) = try destApp t1 with DestKO -> (t1, [||]) in
- let (term2,l2 as appr2) = try destApp t2 with DestKO -> (t2, [||]) in
+ let (term1,l1 as appr1) = try destApp evd t1 with DestKO -> (t1, [||]) in
+ let (term2,l2 as appr2) = try destApp evd t2 with DestKO -> (t2, [||]) in
let () = if !debug_unification then
let open Pp in
Feedback.msg_notice (v 0 (str "Heuristic:" ++ spc () ++ print_constr t1
++ cut () ++ print_constr t2 ++ cut ())) in
let app_empty = Array.is_empty l1 && Array.is_empty l2 in
- match kind_of_term term1, kind_of_term term2 with
+ match EConstr.kind evd term1, EConstr.kind evd term2 with
| Evar (evk1,args1), (Rel _|Var _) when app_empty
- && List.for_all (fun a -> Term.eq_constr a term2 || isEvar a)
+ && List.for_all (fun a -> EConstr.eq_constr evd a term2 || isEvar evd a)
(remove_instance_local_defs evd evk1 args1) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
@@ -1143,9 +1169,9 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
| Some evd -> Success evd
| None ->
let reason = ProblemBeyondCapabilities in
- UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason)))
+ UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason)))
| (Rel _|Var _), Evar (evk2,args2) when app_empty
- && List.for_all (fun a -> Term.eq_constr a term1 || isEvar a)
+ && List.for_all (fun a -> EConstr.eq_constr evd a term1 || isEvar evd a)
(remove_instance_local_defs evd evk2 args2) ->
(* The typical kind of constraint coming from pattern-matching return
type inference *)
@@ -1153,7 +1179,7 @@ let apply_conversion_problem_heuristic ts env evd pbty t1 t2 =
| Some evd -> Success evd
| None ->
let reason = ProblemBeyondCapabilities in
- UnifFailure (evd, CannotSolveConstraint ((pbty,env,t1,t2),reason)))
+ UnifFailure (evd, CannotSolveConstraint (to_pb (pbty,env,t1,t2),reason)))
| Evar (evk1,args1), Evar (evk2,args2) when Evar.equal evk1 evk2 ->
let f env evd pbty x y = is_fconv ~reds:ts pbty env evd x y in
Success (solve_refl ~can_drop:true f env evd
@@ -1193,10 +1219,10 @@ let error_cannot_unify env evd pb ?reason t1 t2 =
let check_problems_are_solved env evd =
match snd (extract_all_conv_pbs evd) with
- | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb t1 t2
+ | (pbty,env,t1,t2) as pb::_ -> error_cannot_unify env evd pb (EConstr.of_constr t1) (EConstr.of_constr t2)
| _ -> ()
-exception MaxUndefined of (Evar.t * evar_info * constr list)
+exception MaxUndefined of (Evar.t * evar_info * Constr.t list)
let max_undefined_with_candidates evd =
let fold evk evi () = match evi.evar_candidates with
@@ -1224,7 +1250,7 @@ let rec solve_unconstrained_evars_with_candidates ts evd =
| a::l ->
try
let conv_algo = evar_conv_x ts in
- let evd = check_evar_instance evd evk a conv_algo in
+ let evd = check_evar_instance evd evk (EConstr.of_constr a) conv_algo in
let evd = Evd.define evk a evd in
match reconsider_unif_constraints conv_algo evd with
| Success evd -> solve_unconstrained_evars_with_candidates ts evd
@@ -1246,7 +1272,7 @@ let solve_unconstrained_impossible_cases env evd =
let ty = j_type j in
let conv_algo = evar_conv_x full_transparent_state in
let evd' = check_evar_instance evd' evk ty conv_algo in
- Evd.define evk ty evd'
+ Evd.define evk (EConstr.Unsafe.to_constr ty) evd'
| _ -> evd') evd evd
let solve_unif_constraints_with_heuristics env
@@ -1255,6 +1281,8 @@ let solve_unif_constraints_with_heuristics env
let rec aux evd pbs progress stuck =
match pbs with
| (pbty,env,t1,t2 as pb) :: pbs ->
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
(match apply_conversion_problem_heuristic ts env evd pbty t1 t2 with
| Success evd' ->
let (evd', rest) = extract_all_conv_pbs evd' in
@@ -1271,6 +1299,8 @@ let solve_unif_constraints_with_heuristics env
match stuck with
| [] -> (* We're finished *) evd
| (pbty,env,t1,t2 as pb) :: _ ->
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
(* There remains stuck problems *)
error_cannot_unify env evd pb t1 t2
in
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index 2231e5bc3..fc07f0fbe 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Environ
open Reductionops
open Evd
@@ -46,18 +47,18 @@ val check_problems_are_solved : env -> evar_map -> unit
(** Check if a canonical structure is applicable *)
val check_conv_record : env -> evar_map ->
- constr * types Stack.t -> constr * types Stack.t ->
+ state -> state ->
Univ.universe_context_set * (constr * constr)
* constr * constr list * (constr Stack.t * constr Stack.t) *
- (constr Stack.t * types Stack.t) *
- (constr Stack.t * types Stack.t) * constr *
+ (constr Stack.t * constr Stack.t) *
+ (constr Stack.t * constr Stack.t) * constr *
(int option * constr)
(** Try to solve problems of the form ?x[args] = c by second-order
matching, using typing to select occurrences *)
val second_order_matching : transparent_state -> env -> evar_map ->
- existential -> occurrences option list -> constr -> evar_map * bool
+ EConstr.existential -> occurrences option list -> constr -> evar_map * bool
(** Declare function to enforce evars resolution by using typing constraints *)
diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml
index 06f619410..c5ae684e3 100644
--- a/pretyping/evardefine.ml
+++ b/pretyping/evardefine.ml
@@ -10,10 +10,11 @@ open Util
open Pp
open Names
open Term
-open Vars
open Termops
-open Namegen
open Environ
+open EConstr
+open Vars
+open Namegen
open Evd
open Evarutil
open Pretype_errors
@@ -27,21 +28,22 @@ let new_evar_unsafe env evd ?src ?filter ?candidates ?store ?naming ?principal t
(Sigma.to_evar_map evd, evk)
let env_nf_evar sigma env =
+ let nf_evar c = nf_evar sigma c in
process_rel_context
- (fun d e -> push_rel (RelDecl.map_constr (nf_evar sigma) d) e) env
+ (fun d e -> push_rel (RelDecl.map_constr nf_evar d) e) env
let env_nf_betaiotaevar sigma env =
process_rel_context
(fun d e ->
- push_rel (RelDecl.map_constr (Reductionops.nf_betaiota sigma) d) e) env
+ push_rel (RelDecl.map_constr (fun c -> Reductionops.nf_betaiota sigma c) d) e) env
(****************************************)
(* Operations on value/type constraints *)
(****************************************)
-type type_constraint = types option
+type type_constraint = EConstr.types option
-type val_constraint = constr option
+type val_constraint = EConstr.constr option
(* Old comment...
* Basically, we have the following kind of constraints (in increasing
@@ -78,8 +80,8 @@ let define_pure_evar_as_product evd evk =
let evi = Evd.find_undefined evd evk in
let evenv = evar_env evi in
let id = next_ident_away idx (ids_of_named_context (evar_context evi)) in
- let concl = Reductionops.whd_all evenv evd evi.evar_concl in
- let s = destSort concl in
+ let concl = Reductionops.whd_all evenv evd (EConstr.of_constr evi.evar_concl) in
+ let s = destSort evd concl in
let evd1,(dom,u1) =
let evd = Sigma.Unsafe.of_evar_map evd in
let Sigma (e, evd1, _) = new_type_evar evenv evd univ_flexible_alg ~filter:(evar_filter evi) in
@@ -89,7 +91,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 s then
+ if is_prop_sort (ESorts.kind evd1 s) then
(* Impredicative product, conclusion must fall in [Prop]. *)
new_evar_unsafe newenv evd1 concl ~src ~filter
else
@@ -100,11 +102,11 @@ let define_pure_evar_as_product evd evk =
(Sigma.to_evar_map evd3, e)
in
let prods = Univ.sup (univ_of_sort u1) (univ_of_sort srng) in
- let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) s in
+ let evd3 = Evd.set_leq_sort evenv evd3 (Type prods) (ESorts.kind evd1 s) in
evd3, rng
in
let prod = mkProd (Name id, dom, subst_var id rng) in
- let evd3 = Evd.define evk prod evd2 in
+ let evd3 = Evd.define evk (EConstr.Unsafe.to_constr prod) evd2 in
evd3,prod
(* Refine an applied evar to a product and returns its instantiation *)
@@ -112,11 +114,11 @@ let define_pure_evar_as_product evd evk =
let define_evar_as_product evd (evk,args) =
let evd,prod = define_pure_evar_as_product evd evk in
(* Quick way to compute the instantiation of evk with args *)
- let na,dom,rng = destProd prod in
- let evdom = mkEvar (fst (destEvar dom), args) in
+ let na,dom,rng = destProd evd prod in
+ let evdom = mkEvar (fst (destEvar evd dom), args) in
let evrngargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evrng = mkEvar (fst (destEvar rng), evrngargs) in
- evd,mkProd (na, evdom, evrng)
+ let evrng = mkEvar (fst (destEvar evd rng), evrngargs) in
+ evd, mkProd (na, evdom, evrng)
(* Refine an evar with an abstraction
@@ -131,10 +133,10 @@ let define_pure_evar_as_lambda env evd evk =
let open Context.Named.Declaration in
let evi = Evd.find_undefined evd evk in
let evenv = evar_env evi in
- let typ = Reductionops.whd_all evenv evd (evar_concl evi) in
- let evd1,(na,dom,rng) = match kind_of_term typ with
+ let typ = Reductionops.whd_all evenv evd (EConstr.of_constr (evar_concl evi)) in
+ let evd1,(na,dom,rng) = match EConstr.kind evd typ with
| Prod (na,dom,rng) -> (evd,(na,dom,rng))
- | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd typ
+ | Evar ev' -> let evd,typ = define_evar_as_product evd ev' in evd,destProd evd typ
| _ -> error_not_product env evd typ in
let avoid = ids_of_named_context (evar_context evi) in
let id =
@@ -144,23 +146,23 @@ let define_pure_evar_as_lambda env evd evk =
let src = evar_source evk evd1 in
let evd2,body = new_evar_unsafe newenv evd1 ~src (subst1 (mkVar id) rng) ~filter in
let lam = mkLambda (Name id, dom, subst_var id body) in
- Evd.define evk lam evd2, lam
+ Evd.define evk (EConstr.Unsafe.to_constr lam) evd2, lam
let define_evar_as_lambda env evd (evk,args) =
let evd,lam = define_pure_evar_as_lambda env evd evk in
(* Quick way to compute the instantiation of evk with args *)
- let na,dom,body = destLambda lam in
+ let na,dom,body = destLambda evd lam in
let evbodyargs = Array.cons (mkRel 1) (Array.map (lift 1) args) in
- let evbody = mkEvar (fst (destEvar body), evbodyargs) in
- evd,mkLambda (na, dom, evbody)
+ let evbody = mkEvar (fst (destEvar evd body), evbodyargs) in
+ evd, mkLambda (na, dom, evbody)
let rec evar_absorb_arguments env evd (evk,args as ev) = function
| [] -> evd,ev
| a::l ->
(* TODO: optimize and avoid introducing intermediate evars *)
let evd,lam = define_pure_evar_as_lambda env evd evk in
- let _,_,body = destLambda lam in
- let evk = fst (destEvar body) in
+ let _,_,body = destLambda evd lam in
+ let evk = fst (destEvar evd body) in
evar_absorb_arguments env evd (evk, Array.cons a args) l
(* Refining an evar to a sort *)
@@ -169,10 +171,10 @@ let define_evar_as_sort env evd (ev,args) =
let evd, u = new_univ_variable univ_rigid evd in
let evi = Evd.find_undefined evd ev in
let s = Type u in
- let concl = Reductionops.whd_all (evar_env evi) evd evi.evar_concl in
- let sort = destSort concl in
- let evd' = Evd.define ev (mkSort s) evd in
- Evd.set_leq_sort env evd' (Type (Univ.super u)) sort, s
+ let concl = Reductionops.whd_all (evar_env evi) evd (EConstr.of_constr evi.evar_concl) in
+ let sort = destSort evd concl in
+ let evd' = Evd.define ev (Constr.mkSort s) evd in
+ Evd.set_leq_sort env evd' (Type (Univ.super u)) (ESorts.kind evd' sort), s
(* Propagation of constraints through application and abstraction:
Given a type constraint on a functional term, returns the type
@@ -182,14 +184,14 @@ let define_evar_as_sort env evd (ev,args) =
let split_tycon loc env evd tycon =
let rec real_split evd c =
let t = Reductionops.whd_all env evd c in
- match kind_of_term t with
+ match EConstr.kind evd t with
| Prod (na,dom,rng) -> evd, (na, dom, rng)
| Evar ev (* ev is undefined because of whd_all *) ->
let (evd',prod) = define_evar_as_product evd ev in
- let (_,dom,rng) = destProd prod in
+ let (_,dom,rng) = destProd evd prod in
evd',(Anonymous, dom, rng)
- | App (c,args) when isEvar c ->
- let (evd',lam) = define_evar_as_lambda env evd (destEvar c) in
+ | App (c,args) when isEvar evd c ->
+ let (evd',lam) = define_evar_as_lambda env evd (destEvar evd c) in
real_split evd' (mkApp (lam,args))
| _ -> error_not_product ~loc env evd c
in
@@ -202,6 +204,6 @@ let split_tycon loc env evd tycon =
let valcon_of_tycon x = x
let lift_tycon n = Option.map (lift n)
-let pr_tycon env = function
+let pr_tycon env sigma = function
None -> str "None"
- | Some t -> Termops.print_constr_env env t
+ | Some t -> Termops.print_constr_env env sigma t
diff --git a/pretyping/evardefine.mli b/pretyping/evardefine.mli
index 07b0e69d9..2f7ac4efb 100644
--- a/pretyping/evardefine.mli
+++ b/pretyping/evardefine.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Evd
open Environ
@@ -42,5 +43,5 @@ val define_evar_as_sort : env -> evar_map -> existential -> evar_map * sorts
(** {6 debug pretty-printer:} *)
-val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
+val pr_tycon : env -> evar_map -> type_constraint -> Pp.std_ppcmds
diff --git a/pretyping/evarsolve.ml b/pretyping/evarsolve.ml
index 92662f07d..77086d046 100644
--- a/pretyping/evarsolve.ml
+++ b/pretyping/evarsolve.ml
@@ -6,14 +6,16 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
open Util
open CErrors
open Names
open Term
-open Vars
open Environ
open Termops
open Evd
+open EConstr
+open Vars
open Namegen
open Retyping
open Reductionops
@@ -22,13 +24,13 @@ open Pretype_errors
open Sigma.Notations
let normalize_evar evd ev =
- match kind_of_term (whd_evar evd (mkEvar ev)) with
+ match EConstr.kind evd (mkEvar ev) with
| Evar (evk,args) -> (evk,args)
| _ -> assert false
-let get_polymorphic_positions f =
+let get_polymorphic_positions sigma f =
let open Declarations in
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Ind (ind, u) | Construct ((ind, _), u) ->
let mib,oib = Global.lookup_inductive ind in
(match oib.mind_arity with
@@ -48,6 +50,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
let modified = ref false in
(* direction: true for fresh universes lower than the existing ones *)
let refresh_sort status ~direction s =
+ let s = ESorts.kind !evdref s in
let s' = evd_comb0 (new_sort_variable status) evdref in
let evd =
if direction then set_leq_sort env !evdref s' s
@@ -56,8 +59,10 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
modified := true; evdref := evd; mkSort s'
in
let rec refresh ~onlyalg status ~direction t =
- match kind_of_term t with
- | Sort (Type u as s) ->
+ match EConstr.kind !evdref t with
+ | Sort s ->
+ begin match ESorts.kind !evdref s with
+ | Type u ->
(match Univ.universe_level u with
| None -> refresh_sort status ~direction s
| Some l ->
@@ -69,29 +74,31 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
if onlyalg && alg then
(evdref := Evd.make_flexible_variable !evdref false l; t)
else t))
- | Sort (Prop Pos as s) when refreshset && not direction ->
+ | Prop Pos when refreshset && not direction ->
(* Cannot make a universe "lower" than "Set",
only refreshing when we want higher universes. *)
refresh_sort status ~direction s
+ | _ -> t
+ end
| Prod (na,u,v) ->
mkProd (na, u, refresh ~onlyalg status ~direction v)
| _ -> t
(** Refresh the types of evars under template polymorphic references *)
and refresh_term_evars onevars top t =
- match kind_of_term (whd_evar !evdref t) with
- | App (f, args) when is_template_polymorphic env f ->
- let pos = get_polymorphic_positions f in
+ match EConstr.kind !evdref t with
+ | App (f, args) when is_template_polymorphic env !evdref f ->
+ let pos = get_polymorphic_positions !evdref f in
refresh_polymorphic_positions args pos
- | App (f, args) when top && isEvar f ->
+ | App (f, args) when top && isEvar !evdref f ->
refresh_term_evars true false f;
Array.iter (refresh_term_evars onevars false) args
| Evar (ev, a) when onevars ->
let evi = Evd.find !evdref ev in
- let ty' = refresh ~onlyalg univ_flexible ~direction:true evi.evar_concl in
+ let ty' = refresh ~onlyalg univ_flexible ~direction:true (EConstr.of_constr evi.evar_concl) in
if !modified then
- evdref := Evd.add !evdref ev {evi with evar_concl = ty'}
+ evdref := Evd.add !evdref ev {evi with evar_concl = EConstr.Unsafe.to_constr ty'}
else ()
- | _ -> Constr.iter (refresh_term_evars onevars false) t
+ | _ -> EConstr.iter !evdref (refresh_term_evars onevars false) t
and refresh_polymorphic_positions args pos =
let rec aux i = function
| Some l :: ls ->
@@ -106,7 +113,7 @@ let refresh_universes ?(status=univ_rigid) ?(onlyalg=false) ?(refreshset=false)
in aux 0 pos
in
let t' =
- if isArity t then
+ if isArity !evdref t then
match pbty with
| None ->
(* No cumulativity needed, but we still need to refresh the algebraics *)
@@ -135,6 +142,8 @@ let test_success conv_algo env evd c c' rhs =
is_success (conv_algo env evd c c' rhs)
let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd =
+ let t1 = EConstr.Unsafe.to_constr t1 in
+ let t2 = EConstr.Unsafe.to_constr t2 in
match pbty with
| Some true -> add_conv_pb ~tail (Reduction.CUMUL,env,t1,t2) evd
| Some false -> add_conv_pb ~tail (Reduction.CUMUL,env,t2,t1) evd
@@ -142,18 +151,18 @@ let add_conv_oriented_pb ?(tail=true) (pbty,env,t1,t2) evd =
(* We retype applications to ensure the universe constraints are collected *)
-exception IllTypedInstance of env * types * types
+exception IllTypedInstance of env * EConstr.types * EConstr.types
let recheck_applications conv_algo env evdref t =
let rec aux env t =
- match kind_of_term t with
+ match EConstr.kind !evdref t with
| App (f, args) ->
let () = aux env f in
let fty = Retyping.get_type_of env !evdref f in
let argsty = Array.map (fun x -> aux env x; Retyping.get_type_of env !evdref x) args in
let rec aux i ty =
if i < Array.length argsty then
- match kind_of_term (whd_all env !evdref ty) with
+ match EConstr.kind !evdref (whd_all env !evdref ty) with
| Prod (na, dom, codom) ->
(match conv_algo env !evdref Reduction.CUMUL argsty.(i) dom with
| Success evd -> evdref := evd;
@@ -164,7 +173,7 @@ let recheck_applications conv_algo env evdref t =
else ()
in aux 0 fty
| _ ->
- iter_constr_with_full_binders (fun d env -> push_rel d env) aux env t
+ iter_with_full_binders !evdref (fun d env -> push_rel d env) aux env t
in aux env t
@@ -193,7 +202,7 @@ let restrict_evar_key evd evk filter candidates =
| None -> evar_filter evi
| Some filter -> filter in
let candidates = match candidates with
- | NoUpdate -> evi.evar_candidates
+ | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates
| UpdateWith c -> Some c in
let sigma = Sigma.Unsafe.of_evar_map evd in
let Sigma (evk, sigma, _) = restrict_evar sigma evk filter candidates in
@@ -226,25 +235,22 @@ open Context.Rel.Declaration
let noccur_evar env evd evk c =
let cache = ref Int.Set.empty (* cache for let-ins *) in
let rec occur_rec check_types (k, env as acc) c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (evk',args' as ev') ->
- (match safe_evar_value evd ev' with
- | Some c -> occur_rec check_types acc c
- | None ->
- if Evar.equal evk evk' then raise Occur
- else (if check_types then
- occur_rec false acc (existential_type evd ev');
- Array.iter (occur_rec check_types acc) args'))
+ if Evar.equal evk evk' then raise Occur
+ else (if check_types then
+ occur_rec false acc (existential_type evd ev');
+ Array.iter (occur_rec check_types acc) args')
| Rel i when i > k ->
if not (Int.Set.mem (i-k) !cache) then
let decl = Environ.lookup_rel i env in
if check_types then
- (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (get_type decl)));
+ (cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr (get_type decl))));
(match decl with
| LocalAssum _ -> ()
- | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i b))
+ | LocalDef (_,b,_) -> cache := Int.Set.add (i-k) !cache; occur_rec false acc (lift i (EConstr.of_constr b)))
| Proj (p,c) -> occur_rec true acc c
- | _ -> iter_constr_with_full_binders (fun rd (k,env) -> (succ k, push_rel rd env))
+ | _ -> iter_with_full_binders evd (fun rd (k,env) -> (succ k, push_rel rd env))
(occur_rec check_types) acc c
in
try occur_rec false (0,env) c; true with Occur -> false
@@ -253,158 +259,209 @@ let noccur_evar env evd evk c =
(* Managing chains of local definitons *)
(***************************************)
+type alias =
+| RelAlias of int
+| VarAlias of Id.t
+
+let of_alias = function
+| RelAlias n -> mkRel n
+| VarAlias id -> mkVar id
+
+let to_alias sigma c = match EConstr.kind sigma c with
+| Rel n -> Some (RelAlias n)
+| Var id -> Some (VarAlias id)
+| _ -> None
+
+let is_alias sigma c alias = match EConstr.kind sigma c, alias with
+| Var id, VarAlias id' -> Id.equal id id'
+| Rel n, RelAlias n' -> Int.equal n n'
+| _ -> false
+
+let eq_alias a b = match a, b with
+| RelAlias n, RelAlias m -> Int.equal m n
+| VarAlias id1, VarAlias id2 -> Id.equal id1 id2
+| _ -> false
+
+type aliasing = EConstr.t option * alias list
+
+let empty_aliasing = None, []
+let make_aliasing c = Some c, []
+let push_alias (alias, l) a = (alias, a :: l)
+let lift_aliasing n (alias, l) =
+ let map a = match a with
+ | VarAlias _ -> a
+ | RelAlias m -> RelAlias (m + n)
+ in
+ (Option.map (fun c -> lift n c) alias, List.map map l)
+
+type aliases = {
+ rel_aliases : aliasing Int.Map.t;
+ var_aliases : aliasing Id.Map.t;
+ (** Only contains [VarAlias] *)
+}
+
(* Expand rels and vars that are bound to other rels or vars so that
dependencies in variables are canonically associated to the most ancient
variable in its family of aliased variables *)
-let compute_var_aliases sign =
+let compute_var_aliases sign sigma =
let open Context.Named.Declaration in
List.fold_right (fun decl aliases ->
let id = get_id decl in
match decl with
| LocalDef (_,t,_) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_id =
- try Id.Map.find id' aliases with Not_found -> [] in
- Id.Map.add id (aliases_of_id@[t]) aliases
+ try Id.Map.find id' aliases with Not_found -> empty_aliasing in
+ Id.Map.add id (push_alias aliases_of_id (VarAlias id')) aliases
| _ ->
- Id.Map.add id [t] aliases)
+ Id.Map.add id (make_aliasing t) aliases)
| LocalAssum _ -> aliases)
sign Id.Map.empty
-let compute_rel_aliases var_aliases rels =
+let compute_rel_aliases var_aliases rels sigma =
snd (List.fold_right
(fun decl (n,aliases) ->
(n-1,
match decl with
| LocalDef (_,t,u) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_n =
- try Id.Map.find id' var_aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[t]) aliases
+ try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in
+ Int.Map.add n (push_alias aliases_of_n (VarAlias id')) aliases
| Rel p ->
let aliases_of_n =
- try Int.Map.find (p+n) aliases with Not_found -> [] in
- Int.Map.add n (aliases_of_n@[mkRel (p+n)]) aliases
+ try Int.Map.find (p+n) aliases with Not_found -> empty_aliasing in
+ Int.Map.add n (push_alias aliases_of_n (RelAlias (p+n))) aliases
| _ ->
- Int.Map.add n [lift n (mkCast(t,DEFAULTcast,u))] aliases)
+ Int.Map.add n (make_aliasing (lift n (mkCast(t,DEFAULTcast,u)))) aliases)
| LocalAssum _ -> aliases)
)
rels
(List.length rels,Int.Map.empty))
-let make_alias_map env =
+let make_alias_map env sigma =
(* We compute the chain of aliases for each var and rel *)
- let var_aliases = compute_var_aliases (named_context env) in
- let rel_aliases = compute_rel_aliases var_aliases (rel_context env) in
- (var_aliases,rel_aliases)
+ let var_aliases = compute_var_aliases (named_context env) sigma in
+ let rel_aliases = compute_rel_aliases var_aliases (rel_context env) sigma in
+ { var_aliases; rel_aliases }
-let lift_aliases n (var_aliases,rel_aliases as aliases) =
+let lift_aliases n aliases =
if Int.equal n 0 then aliases else
- (var_aliases,
- Int.Map.fold (fun p l -> Int.Map.add (p+n) (List.map (lift n) l))
- rel_aliases Int.Map.empty)
-
-let get_alias_chain_of aliases x = match kind_of_term x with
- | Rel n -> (try Int.Map.find n (snd aliases) with Not_found -> [])
- | Var id -> (try Id.Map.find id (fst aliases) with Not_found -> [])
- | _ -> []
-
-let normalize_alias_opt aliases x =
- match get_alias_chain_of aliases x with
- | [] -> None
- | a::_ when isRel a || isVar a -> Some a
- | [_] -> None
- | _::a::_ -> Some a
-
-let normalize_alias aliases x =
- match normalize_alias_opt aliases x with
+ let rel_aliases =
+ Int.Map.fold (fun p l -> Int.Map.add (p+n) (lift_aliasing n l))
+ aliases.rel_aliases Int.Map.empty
+ in
+ { aliases with rel_aliases }
+
+let get_alias_chain_of sigma aliases x = match x with
+ | RelAlias n -> (try Int.Map.find n aliases.rel_aliases with Not_found -> empty_aliasing)
+ | VarAlias id -> (try Id.Map.find id aliases.var_aliases with Not_found -> empty_aliasing)
+
+let normalize_alias_opt_alias sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | _, [] -> None
+ | _, a :: _ -> Some a
+
+let normalize_alias_opt sigma aliases x = match to_alias sigma x with
+| None -> None
+| Some a -> normalize_alias_opt_alias sigma aliases a
+
+let normalize_alias sigma aliases x =
+ match normalize_alias_opt_alias sigma aliases x with
| Some a -> a
| None -> x
-let normalize_alias_var var_aliases id =
- destVar (normalize_alias (var_aliases,Int.Map.empty) (mkVar id))
+let normalize_alias_var sigma var_aliases id =
+ let aliases = { var_aliases; rel_aliases = Int.Map.empty } in
+ match normalize_alias sigma aliases (VarAlias id) with
+ | VarAlias id -> id
+ | RelAlias _ -> assert false (** var only aliases to variables *)
-let extend_alias decl (var_aliases,rel_aliases) =
+let extend_alias sigma decl { var_aliases; rel_aliases } =
let rel_aliases =
- Int.Map.fold (fun n l -> Int.Map.add (n+1) (List.map (lift 1) l))
+ Int.Map.fold (fun n l -> Int.Map.add (n+1) (lift_aliasing 1 l))
rel_aliases Int.Map.empty in
let rel_aliases =
match decl with
| LocalDef(_,t,_) ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| Var id' ->
let aliases_of_binder =
- try Id.Map.find id' var_aliases with Not_found -> [] in
- Int.Map.add 1 (aliases_of_binder@[t]) rel_aliases
+ try Id.Map.find id' var_aliases with Not_found -> empty_aliasing in
+ Int.Map.add 1 (push_alias aliases_of_binder (VarAlias id')) rel_aliases
| Rel p ->
let aliases_of_binder =
- try Int.Map.find (p+1) rel_aliases with Not_found -> [] in
- Int.Map.add 1 (aliases_of_binder@[mkRel (p+1)]) rel_aliases
+ try Int.Map.find (p+1) rel_aliases with Not_found -> empty_aliasing in
+ Int.Map.add 1 (push_alias aliases_of_binder (RelAlias (p+1))) rel_aliases
| _ ->
- Int.Map.add 1 [lift 1 t] rel_aliases)
+ Int.Map.add 1 (make_aliasing (lift 1 t)) rel_aliases)
| LocalAssum _ -> rel_aliases in
- (var_aliases, rel_aliases)
-
-let expand_alias_once aliases x =
- match get_alias_chain_of aliases x with
- | [] -> None
- | l -> Some (List.last l)
-
-let expansions_of_var aliases x =
- match get_alias_chain_of aliases x with
- | [] -> [x]
- | a::_ as l when isRel a || isVar a -> x :: List.rev l
- | _::l -> x :: List.rev l
-
-let expansion_of_var aliases x =
- match get_alias_chain_of aliases x with
- | [] -> x
- | a::_ -> a
-
-let rec expand_vars_in_term_using aliases t = match kind_of_term t with
- | Rel _ | Var _ ->
- normalize_alias aliases t
+ { var_aliases; rel_aliases }
+
+let expand_alias_once sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | None, [] -> None
+ | Some a, [] -> Some a
+ | _, l -> Some (of_alias (List.last l))
+
+let expansions_of_var sigma aliases x =
+ let (_, l) = get_alias_chain_of sigma aliases x in
+ x :: List.rev l
+
+let expansion_of_var sigma aliases x =
+ match get_alias_chain_of sigma aliases x with
+ | None, [] -> (false, of_alias x)
+ | Some a, _ -> (true, a)
+ | None, a :: _ -> (true, of_alias a)
+
+let rec expand_vars_in_term_using sigma aliases t = match EConstr.kind sigma t with
+ | Rel n -> of_alias (normalize_alias sigma aliases (RelAlias n))
+ | Var id -> of_alias (normalize_alias sigma aliases (VarAlias id))
| _ ->
- map_constr_with_full_binders
- extend_alias expand_vars_in_term_using aliases t
+ let self aliases c = expand_vars_in_term_using sigma aliases c in
+ map_constr_with_full_binders sigma (extend_alias sigma) self aliases t
-let expand_vars_in_term env = expand_vars_in_term_using (make_alias_map env)
+let expand_vars_in_term env sigma = expand_vars_in_term_using sigma (make_alias_map env sigma)
-let free_vars_and_rels_up_alias_expansion aliases c =
+let free_vars_and_rels_up_alias_expansion sigma aliases c =
let acc1 = ref Int.Set.empty and acc2 = ref Id.Set.empty in
let acc3 = ref Int.Set.empty and acc4 = ref Id.Set.empty in
let cache_rel = ref Int.Set.empty and cache_var = ref Id.Set.empty in
let is_in_cache depth = function
- | Rel n -> Int.Set.mem (n-depth) !cache_rel
- | Var s -> Id.Set.mem s !cache_var
- | _ -> false in
+ | RelAlias n -> Int.Set.mem (n-depth) !cache_rel
+ | VarAlias s -> Id.Set.mem s !cache_var
+ in
let put_in_cache depth = function
- | Rel n -> cache_rel := Int.Set.add (n-depth) !cache_rel
- | Var s -> cache_var := Id.Set.add s !cache_var
- | _ -> () in
+ | RelAlias n -> cache_rel := Int.Set.add (n-depth) !cache_rel
+ | VarAlias s -> cache_var := Id.Set.add s !cache_var
+ in
let rec frec (aliases,depth) c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Rel _ | Var _ as ck ->
+ let ck = match ck with
+ | Rel n -> RelAlias n
+ | Var id -> VarAlias id
+ | _ -> assert false
+ in
if is_in_cache depth ck then () else begin
put_in_cache depth ck;
- let c' = expansion_of_var aliases c in
- (if c != c' then (* expansion, hence a let-in *)
- match kind_of_term c with
- | Var id -> acc4 := Id.Set.add id !acc4
- | Rel n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3
- | _ -> ());
- match kind_of_term c' with
+ let expanded, c' = expansion_of_var sigma aliases ck in
+ (if expanded then (* expansion, hence a let-in *)
+ match ck with
+ | VarAlias id -> acc4 := Id.Set.add id !acc4
+ | RelAlias n -> if n >= depth+1 then acc3 := Int.Set.add (n-depth) !acc3);
+ match EConstr.kind sigma c' with
| Var id -> acc2 := Id.Set.add id !acc2
| Rel n -> if n >= depth+1 then acc1 := Int.Set.add (n-depth) !acc1
| _ -> frec (aliases,depth) c end
| Const _ | Ind _ | Construct _ ->
- acc2 := Id.Set.union (vars_of_global (Global.env()) c) !acc2
+ acc2 := Id.Set.union (vars_of_global (Global.env()) (EConstr.to_constr sigma c)) !acc2
| _ ->
- iter_constr_with_full_binders
- (fun d (aliases,depth) -> (extend_alias d aliases,depth+1))
+ iter_with_full_binders sigma
+ (fun d (aliases,depth) -> (extend_alias sigma d aliases,depth+1))
frec (aliases,depth) c
in
frec (aliases,0) c;
@@ -414,42 +471,45 @@ let free_vars_and_rels_up_alias_expansion aliases c =
(* Managing pattern-unification *)
(********************************)
-let rec expand_and_check_vars aliases = function
+let map_all f l =
+ let rec map_aux f l = match l with
| [] -> []
- | a::l when isRel a || isVar a ->
- let a = expansion_of_var aliases a in
- if isRel a || isVar a then a :: expand_and_check_vars aliases l
- else raise Exit
- | _ ->
- raise Exit
-
-module Constrhash = Hashtbl.Make
- (struct type t = constr
- let equal = Term.eq_constr
- let hash = hash_constr
- end)
-
-let constr_list_distinct l =
- let visited = Constrhash.create 23 in
- let rec loop = function
- | h::t ->
- if Constrhash.mem visited h then false
- else (Constrhash.add visited h h; loop t)
- | [] -> true
- in loop l
-
-let get_actual_deps aliases l t =
- if occur_meta_or_existential t then
+ | x :: l ->
+ match f x with
+ | None -> raise Exit
+ | Some y -> y :: map_aux f l
+ in
+ try Some (map_aux f l) with Exit -> None
+
+let expand_and_check_vars sigma aliases l =
+ let map a = match get_alias_chain_of sigma aliases a with
+ | None, [] -> Some a
+ | None, a :: _ -> Some a
+ | Some _, _ -> None
+ in
+ map_all map l
+
+let alias_distinct l =
+ let rec check (rels, vars) = function
+ | [] -> true
+ | RelAlias n :: l ->
+ not (Int.Set.mem n rels) && check (Int.Set.add n rels, vars) l
+ | VarAlias id :: l ->
+ not (Id.Set.mem id vars) && check (rels, Id.Set.add id vars) l
+ in
+ check (Int.Set.empty, Id.Set.empty) l
+
+let get_actual_deps evd aliases l t =
+ if occur_meta_or_existential evd t then
(* Probably no restrictions on allowed vars in presence of evars *)
l
else
(* Probably strong restrictions coming from t being evar-closed *)
- let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion aliases t in
- List.filter (fun c ->
- match kind_of_term c with
- | Var id -> Id.Set.mem id fv_ids
- | Rel n -> Int.Set.mem n fv_rels
- | _ -> assert false) l
+ let (fv_rels,fv_ids,_,_) = free_vars_and_rels_up_alias_expansion evd aliases t in
+ List.filter (function
+ | VarAlias id -> Id.Set.mem id fv_ids
+ | RelAlias n -> Int.Set.mem n fv_rels
+ ) l
open Context.Named.Declaration
let remove_instance_local_defs evd evk args =
@@ -468,35 +528,42 @@ let remove_instance_local_defs evd evk args =
(* Check if an applied evar "?X[args] l" is a Miller's pattern *)
-let find_unification_pattern_args env l t =
- if List.for_all (fun x -> isRel x || isVar x) l (* common failure case *) then
- let aliases = make_alias_map env in
- match (try Some (expand_and_check_vars aliases l) with Exit -> None) with
- | Some l as x when constr_list_distinct (get_actual_deps aliases l t) -> x
- | _ -> None
- else
- None
+let find_unification_pattern_args env evd l t =
+ let aliases = make_alias_map env evd in
+ match expand_and_check_vars evd aliases l with
+ | Some l as x when alias_distinct (get_actual_deps evd aliases l t) -> x
+ | _ -> None
-let is_unification_pattern_meta env nb m l t =
+let is_unification_pattern_meta env evd nb m l t =
(* Variables from context and rels > nb are implicitly all there *)
(* so we need to be a rel <= nb *)
- if List.for_all (fun x -> isRel x && destRel x <= nb) l then
- match find_unification_pattern_args env l t with
- | Some _ as x when not (dependent (mkMeta m) t) -> x
+ let map a = match EConstr.kind evd a with
+ | Rel n -> if n <= nb then Some (RelAlias n) else None
+ | _ -> None
+ in
+ match map_all map l with
+ | Some l ->
+ begin match find_unification_pattern_args env evd l t with
+ | Some _ as x when not (dependent evd (mkMeta m) t) -> x
| _ -> None
- else
+ end
+ | None ->
None
let is_unification_pattern_evar env evd (evk,args) l t =
- if List.for_all (fun x -> isRel x || isVar x) l
- && noccur_evar env evd evk t
- then
+ match map_all (fun c -> to_alias evd c) l with
+ | Some l when noccur_evar env evd evk t ->
let args = remove_instance_local_defs evd evk args in
+ let args = map_all (fun c -> to_alias evd c) args in
+ begin match args with
+ | None -> None
+ | Some args ->
let n = List.length args in
- match find_unification_pattern_args env (args @ l) t with
+ match find_unification_pattern_args env evd (args @ l) t with
| Some l -> Some (List.skipn n l)
| _ -> None
- else None
+ end
+ | _ -> None
let is_unification_pattern_pure_evar env evd (evk,args) t =
let is_ev = is_unification_pattern_evar env evd (evk,args) [] t in
@@ -505,8 +572,8 @@ let is_unification_pattern_pure_evar env evd (evk,args) t =
| Some _ -> true
let is_unification_pattern (env,nb) evd f l t =
- match kind_of_term f with
- | Meta m -> is_unification_pattern_meta env nb m l t
+ match EConstr.kind evd f with
+ | Meta m -> is_unification_pattern_meta env evd nb m l t
| Evar ev -> is_unification_pattern_evar env evd ev l t
| _ -> None
@@ -517,18 +584,18 @@ let is_unification_pattern (env,nb) evd f l t =
*implicitly* depend on Vars but lambda abstraction will not reflect this
dependency: ?X x = ?1 (?1 is a meta) will return \_.?1 while it should
return \y. ?1{x\y} (non constant function if ?1 depends on x) (BB) *)
-let solve_pattern_eqn env l c =
+let solve_pattern_eqn env sigma l c =
let c' = List.fold_right (fun a c ->
- let c' = subst_term (lift 1 a) (lift 1 c) in
- match kind_of_term a with
+ let c' = subst_term sigma (lift 1 (of_alias a)) (lift 1 c) in
+ match a with
(* Rem: if [a] links to a let-in, do as if it were an assumption *)
- | Rel n ->
+ | RelAlias n ->
let open Context.Rel.Declaration in
let d = map_constr (lift n) (lookup_rel n env) in
mkLambda_or_LetIn d c'
- | Var id ->
+ | VarAlias id ->
let d = lookup_named id env in mkNamedLambda_or_LetIn d c'
- | _ -> assert false)
+ )
l c in
(* Warning: we may miss some opportunity to eta-reduce more since c'
is not in normal form *)
@@ -550,35 +617,34 @@ let solve_pattern_eqn env l c =
let make_projectable_subst aliases sigma evi args =
let sign = evar_filtered_context evi in
- let evar_aliases = compute_var_aliases sign in
+ let sign = List.map (fun d -> map_named_decl EConstr.of_constr d) sign in
+ let evar_aliases = compute_var_aliases sign sigma in
let (_,full_subst,cstr_subst) =
List.fold_right
(fun decl (args,all,cstrs) ->
match decl,args with
| LocalAssum (id,c), a::rest ->
- let a = whd_evar sigma a in
let cstrs =
- let a',args = decompose_app_vect a in
- match kind_of_term a' with
+ let a',args = decompose_app_vect sigma a in
+ match EConstr.kind sigma a' with
| Construct cstr ->
let l = try Constrmap.find (fst cstr) cstrs with Not_found -> [] in
Constrmap.add (fst cstr) ((args,id)::l) cstrs
| _ -> cstrs in
- (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs)
+ (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs)
| LocalDef (id,c,_), a::rest ->
- let a = whd_evar sigma a in
- (match kind_of_term c with
+ (match EConstr.kind sigma c with
| Var id' ->
- let idc = normalize_alias_var evar_aliases id' in
+ let idc = normalize_alias_var sigma evar_aliases id' in
let sub = try Id.Map.find idc all with Not_found -> [] in
- if List.exists (fun (c,_,_) -> Term.eq_constr a c) sub then
+ if List.exists (fun (c,_,_) -> EConstr.eq_constr sigma a c) sub then
(rest,all,cstrs)
else
(rest,
- Id.Map.add idc ((a,normalize_alias_opt aliases a,id)::sub) all,
+ Id.Map.add idc ((a,normalize_alias_opt sigma aliases a,id)::sub) all,
cstrs)
| _ ->
- (rest,Id.Map.add id [a,normalize_alias_opt aliases a,id] all,cstrs))
+ (rest,Id.Map.add id [a,normalize_alias_opt sigma aliases a,id] all,cstrs))
| _ -> anomaly (Pp.str "Instance does not match its signature"))
sign (Array.rev_to_list args,Id.Map.empty,Constrmap.empty) in
(full_subst,cstr_subst)
@@ -598,10 +664,11 @@ let define_evar_from_virtual_equation define_fun env evd src t_in_env ty_t_in_si
let Sigma (evar_in_env, evd, _) = new_evar_instance sign evd ty_t_in_sign ~filter ~src inst_in_env in
let evd = Sigma.to_evar_map evd in
let t_in_env = whd_evar evd t_in_env in
- let evd = define_fun env evd None (destEvar evar_in_env) t_in_env in
+ let (evk, _) = destEvar evd evar_in_env in
+ let evd = define_fun env evd None (destEvar evd evar_in_env) t_in_env in
let ctxt = named_context_of_val sign in
let inst_in_sign = inst_of_vars (Filter.filter_list filter ctxt) in
- let evar_in_sign = mkEvar (fst (destEvar evar_in_env), inst_in_sign) in
+ let evar_in_sign = mkEvar (evk, inst_in_sign) in
(evd,whd_evar evd evar_in_sign)
(* We have x1..xq |- ?e1 : τ and had to solve something like
@@ -627,7 +694,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
if Evd.is_defined evd evk1 then
(* Some circularity somewhere (see e.g. #3209) *)
raise MorePreciseOccurCheckNeeeded;
- let (evk1,args1) = destEvar (whd_evar evd (mkEvar (evk1,args1))) in
+ let (evk1,args1) = destEvar evd (mkEvar (evk1,args1)) in
let evi1 = Evd.find_undefined evd evk1 in
let env1,rel_sign = env_rel_context_chop k env in
let sign1 = evar_hyps evi1 in
@@ -669,7 +736,7 @@ let materialize_evar define_fun env evd k (evk1,args1) ty_in_env =
let Sigma (ev2_in_sign, evd, _) =
new_evar_instance sign2 evd ev2ty_in_sign ~filter:filter2 ~src inst2_in_sign in
let evd = Sigma.to_evar_map evd in
- let ev2_in_env = (fst (destEvar ev2_in_sign), Array.of_list inst2_in_env) in
+ let ev2_in_env = (fst (destEvar evd ev2_in_sign), Array.of_list inst2_in_env) in
(evd, ev2_in_sign, ev2_in_env)
let restrict_upon_filter evd evk p args =
@@ -690,7 +757,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst =
List.filter (fun (args',id) ->
(* is_conv is maybe too strong (and source of useless computation) *)
(* (at least expansion of aliases is needed) *)
- Array.for_all2 (is_conv env evd) args args') l in
+ Array.for_all2 (fun c1 c2 -> is_conv env evd c1 c2) args args') l in
List.map snd l
with Not_found ->
[]
@@ -729,7 +796,7 @@ let find_projectable_constructor env evd cstr k args cstr_subst =
type evar_projection =
| ProjectVar
-| ProjectEvar of existential * evar_info * Id.t * evar_projection
+| ProjectEvar of EConstr.existential * evar_info * Id.t * evar_projection
exception NotUnique
exception NotUniqueInType of (Id.t * evar_projection) list
@@ -737,19 +804,18 @@ exception NotUniqueInType of (Id.t * evar_projection) list
let rec assoc_up_to_alias sigma aliases y yc = function
| [] -> raise Not_found
| (c,cc,id)::l ->
- let c' = whd_evar sigma c in
- if Term.eq_constr y c' then id
+ if is_alias sigma c y then id
else
match l with
| _ :: _ -> assoc_up_to_alias sigma aliases y yc l
| [] ->
(* Last chance, we reason up to alias conversion *)
- match (if c == c' then cc else normalize_alias_opt aliases c') with
- | Some cc when Term.eq_constr yc cc -> id
- | _ -> if Term.eq_constr yc c then id else raise Not_found
+ match (normalize_alias_opt sigma aliases c) with
+ | Some cc when eq_alias yc cc -> id
+ | _ -> if is_alias sigma c yc then id else raise Not_found
let rec find_projectable_vars with_evars aliases sigma y subst =
- let yc = normalize_alias aliases y in
+ let yc = normalize_alias sigma aliases y in
let is_projectable idc idcl subst' =
(* First test if some [id] aliased to [idc] is bound to [y] in [subst] *)
try
@@ -759,12 +825,12 @@ let rec find_projectable_vars with_evars aliases sigma y subst =
(* Then test if [idc] is (indirectly) bound in [subst] to some evar *)
(* projectable on [y] *)
if with_evars then
- let f (c,_,id) = isEvar c && is_undefined sigma (fst (destEvar c)) in
+ let f (c,_,id) = isEvar sigma c in
let idcl' = List.filter f idcl in
match idcl' with
| [c,_,id] ->
begin
- let (evk,argsv as t) = destEvar c in
+ let (evk,argsv as t) = destEvar sigma c in
let evi = Evd.find sigma evk in
let subst,_ = make_projectable_subst aliases sigma evi argsv in
let l = find_projectable_vars with_evars aliases sigma y subst in
@@ -812,19 +878,18 @@ let rec find_solution_type evarenv = function
let rec do_projection_effects define_fun env ty evd = function
| ProjectVar -> evd
| ProjectEvar ((evk,argsv),evi,id,p) ->
- let evd = Evd.define evk (mkVar id) evd in
+ let evd = Evd.define evk (Constr.mkVar id) evd in
(* TODO: simplify constraints involving evk *)
let evd = do_projection_effects define_fun env ty evd p in
let ty = whd_all env evd (Lazy.force ty) in
- if not (isSort ty) then
+ if not (isSort evd ty) then
(* Don't try to instantiate if a sort because if evar_concl is an
evar it may commit to a univ level which is not the right
one (however, regarding coercions, because t is obtained by
unif, we know that no coercion can be inserted) *)
let subst = make_pure_subst evi argsv in
- let ty' = replace_vars subst evi.evar_concl in
- let ty' = whd_evar evd ty' in
- if isEvar ty' then define_fun env evd (Some false) (destEvar ty') ty else evd
+ let ty' = replace_vars subst (EConstr.of_constr evi.evar_concl) in
+ if isEvar evd ty' then define_fun env evd (Some false) (destEvar evd ty') ty else evd
else
evd
@@ -850,7 +915,7 @@ let rec do_projection_effects define_fun env ty evd = function
type projectibility_kind =
| NoUniqueProjection
- | UniqueProjection of constr * evar_projection list
+ | UniqueProjection of EConstr.constr * evar_projection list
type projectibility_status =
| CannotInvert
@@ -859,15 +924,14 @@ type projectibility_status =
let invert_arg_from_subst evd aliases k0 subst_in_env_extended_with_k_binders c_in_env_extended_with_k_binders =
let effects = ref [] in
let rec aux k t =
- let t = whd_evar evd t in
- match kind_of_term t with
- | Rel i when i>k0+k -> aux' k (mkRel (i-k))
- | Var id -> aux' k t
- | _ -> map_constr_with_binders succ aux k t
+ match EConstr.kind evd t with
+ | Rel i when i>k0+k -> aux' k (RelAlias (i-k))
+ | Var id -> aux' k (VarAlias id)
+ | _ -> map_with_binders evd succ aux k t
and aux' k t =
try project_with_effects aliases evd effects t subst_in_env_extended_with_k_binders
with Not_found ->
- match expand_alias_once aliases t with
+ match expand_alias_once evd aliases t with
| None -> raise Not_found
| Some c -> aux k c in
try
@@ -931,23 +995,23 @@ let invert_invertible_arg fullenv evd aliases k (evk,argsv) args' =
let set_of_evctx l =
List.fold_left (fun s decl -> Id.Set.add (get_id decl) s) Id.Set.empty l
-let filter_effective_candidates evi filter candidates =
+let filter_effective_candidates evd evi filter candidates =
match filter with
| None -> candidates
| Some filter ->
let ids = set_of_evctx (Filter.filter_list filter (evar_context evi)) in
- List.filter (fun a -> Id.Set.subset (collect_vars a) ids) candidates
+ List.filter (fun a -> Id.Set.subset (collect_vars evd a) ids) candidates
let filter_candidates evd evk filter candidates_update =
let evi = Evd.find_undefined evd evk in
let candidates = match candidates_update with
- | NoUpdate -> evi.evar_candidates
+ | NoUpdate -> Option.map (fun l -> List.map EConstr.of_constr l) evi.evar_candidates
| UpdateWith c -> Some c
in
match candidates with
| None -> NoUpdate
| Some l ->
- let l' = filter_effective_candidates evi filter l in
+ let l' = filter_effective_candidates evd evi filter l in
if List.length l = List.length l' && candidates_update = NoUpdate then
NoUpdate
else
@@ -960,13 +1024,14 @@ let closure_of_filter evd evk = function
| None -> None
| Some filter ->
let evi = Evd.find_undefined evd evk in
- let vars = collect_vars (Evarutil.nf_evar evd (evar_concl evi)) in
+ let vars = collect_vars evd (EConstr.of_constr (evar_concl evi)) in
let test b decl = b || Idset.mem (get_id decl) vars ||
match decl with
| LocalAssum _ ->
false
| LocalDef (_,c,_) ->
- not (isRel c || isVar c)
+ let c = EConstr.of_constr c in
+ not (isRel evd c || isVar evd c)
in
let newfilter = Filter.map_along test filter (evar_context evi) in
(* Now ensure that restriction is at least what is was originally *)
@@ -989,7 +1054,7 @@ let restrict_hyps evd evk filter candidates =
let typablefilter = closure_of_filter evd evk (Some filter) in
(typablefilter,candidates)
-exception EvarSolvedWhileRestricting of evar_map * constr
+exception EvarSolvedWhileRestricting of evar_map * EConstr.constr
let do_restrict_hyps evd (evk,args as ev) filter candidates =
let filter,candidates = match filter with
@@ -998,8 +1063,8 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates =
match candidates,filter with
| UpdateWith [], _ -> error "Not solvable."
| UpdateWith [nc],_ ->
- let evd = Evd.define evk nc evd in
- raise (EvarSolvedWhileRestricting (evd,whd_evar evd (mkEvar ev)))
+ let evd = Evd.define evk (EConstr.Unsafe.to_constr nc) evd in
+ raise (EvarSolvedWhileRestricting (evd,mkEvar ev))
| NoUpdate, None -> evd,ev
| _ -> restrict_applied_evar evd ev filter candidates
@@ -1007,7 +1072,7 @@ let do_restrict_hyps evd (evk,args as ev) filter candidates =
(* ?e is assumed to have no candidates *)
let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
- let rhs = expand_vars_in_term env rhs in
+ let rhs = expand_vars_in_term env evd rhs in
let filter =
restrict_upon_filter evd evk
(* Keep only variables that occur in rhs *)
@@ -1017,8 +1082,8 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
(* that says that the body is hidden. Note that expand_vars_in_term *)
(* expands only rels and vars aliases, not rels or vars bound to an *)
(* arbitrary complex term *)
- (fun a -> not (isRel a || isVar a)
- || dependent a rhs || List.exists (fun (id,_) -> isVarId id a) sols)
+ (fun a -> not (isRel evd a || isVar evd a)
+ || dependent evd a rhs || List.exists (fun (id,_) -> isVarId evd id a) sols)
argsv in
let filter = closure_of_filter evd evk filter in
let candidates = extract_candidates sols in
@@ -1049,6 +1114,9 @@ let postpone_non_unique_projection env evd pbty (evk,argsv as ev) sols rhs =
* Note: argument f is the function used to instantiate evars.
*)
+let instantiate_evar_array evi c args =
+ EConstr.of_constr (instantiate_evar_array evi (EConstr.Unsafe.to_constr c) (Array.map EConstr.Unsafe.to_constr args))
+
let filter_compatible_candidates conv_algo env evd evi args rhs c =
let c' = instantiate_evar_array evi c args in
match conv_algo env evd Reduction.CONV rhs c' with
@@ -1068,7 +1136,9 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
| _, None -> filter_candidates evd evk1 filter1 NoUpdate
| None, Some _ -> raise DoesNotPreserveCandidateRestriction
| Some l1, Some l2 ->
- let l1 = filter_effective_candidates evi1 filter1 l1 in
+ let l1 = List.map EConstr.of_constr l1 in
+ let l2 = List.map EConstr.of_constr l2 in
+ let l1 = filter_effective_candidates evd evi1 filter1 l1 in
let l1' = List.filter (fun c1 ->
let c1' = instantiate_evar_array evi1 c1 argsv1 in
let filter c2 =
@@ -1082,7 +1152,7 @@ let restrict_candidates conv_algo env evd filter1 (evk1,argsv1) (evk2,argsv2) =
if Int.equal (List.length l1) (List.length l1') then NoUpdate
else UpdateWith l1'
-exception CannotProject of evar_map * existential
+exception CannotProject of evar_map * EConstr.existential
(* Assume that FV(?n[x1:=t1..xn:=tn]) belongs to some set U.
Can ?n be instantiated by a term u depending essentially on xi such that the
@@ -1099,10 +1169,8 @@ exception CannotProject of evar_map * existential
*)
let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t =
- let f,args2 = decompose_app_vect t in
- let f,args1 = decompose_app_vect (whd_evar evd f) in
- let args = Array.append args1 args2 in
- match kind_of_term f with
+ let f,args = decompose_app_vect evd t in
+ match EConstr.kind evd f with
| Construct ((ind,_),u) ->
let n = Inductiveops.inductive_nparams ind in
if n > Array.length args then true (* We don't try to be more clever *)
@@ -1118,30 +1186,32 @@ let rec is_constrainable_in top evd k (ev,(fv_rels,fv_ids) as g) t =
| _ -> (* We don't try to be more clever *) true
let has_constrainable_free_vars env evd aliases force k ev (fv_rels,fv_ids,let_rels,let_ids) t =
- let t' = expansion_of_var aliases t in
- if t' != t then
+ match to_alias evd t with
+ | Some t ->
+ let expanded, t' = expansion_of_var evd aliases t in
+ if expanded then
(* t is a local definition, we keep it only if appears in the list *)
(* of let-in variables effectively occurring on the right-hand side, *)
(* which is the only reason to keep it when inverting arguments *)
- match kind_of_term t with
- | Var id -> Id.Set.mem id let_ids
- | Rel n -> Int.Set.mem n let_rels
- | _ -> assert false
- else
+ match t with
+ | VarAlias id -> Id.Set.mem id let_ids
+ | RelAlias n -> Int.Set.mem n let_rels
+ else begin match t with
+ | VarAlias id -> Id.Set.mem id fv_ids
+ | RelAlias n -> n <= k || Int.Set.mem n fv_rels
+ end
+ | None ->
(* t is an instance for a proper variable; we filter it along *)
(* the free variables allowed to occur *)
- match kind_of_term t with
- | Var id -> Id.Set.mem id fv_ids
- | Rel n -> n <= k || Int.Set.mem n fv_rels
- | _ -> (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t
+ (not force || noccur_evar env evd ev t) && is_constrainable_in true evd k (ev,(fv_rels,fv_ids)) t
-exception EvarSolvedOnTheFly of evar_map * constr
+exception EvarSolvedOnTheFly of evar_map * EConstr.constr
(* Try to project evk1[argsv1] on evk2[argsv2], if [ev1] is a pattern on
the common domain of definition *)
let project_evar_on_evar force g env evd aliases k2 pbty (evk1,argsv1 as ev1) (evk2,argsv2 as ev2) =
(* Apply filtering on ev1 so that fvs(ev1) are in fvs(ev2). *)
- let fvs2 = free_vars_and_rels_up_alias_expansion aliases (mkEvar ev2) in
+ let fvs2 = free_vars_and_rels_up_alias_expansion evd aliases (mkEvar ev2) in
let filter1 = restrict_upon_filter evd evk1
(has_constrainable_free_vars env evd aliases force k2 evk2 fvs2)
argsv1 in
@@ -1173,9 +1243,9 @@ let check_evar_instance evd evk1 body conv_algo =
try Retyping.get_type_of ~lax:true evenv evd body
with Retyping.RetypeError _ -> error "Ill-typed evar instance"
in
- match conv_algo evenv evd Reduction.CUMUL ty evi.evar_concl with
+ match conv_algo evenv evd Reduction.CUMUL ty (EConstr.of_constr evi.evar_concl) with
| Success evd -> evd
- | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,evi.evar_concl))
+ | UnifFailure _ -> raise (IllTypedInstance (evenv,ty,EConstr.of_constr evi.evar_concl))
let update_evar_source ev1 ev2 evd =
let loc, evs2 = evar_source ev2 evd in
@@ -1188,8 +1258,8 @@ let update_evar_source ev1 ev2 evd =
let solve_evar_evar_l2r force f g env evd aliases pbty ev1 (evk2,_ as ev2) =
try
let evd,body = project_evar_on_evar force g env evd aliases 0 pbty ev1 ev2 in
- let evd' = Evd.define evk2 body evd in
- let evd' = update_evar_source (fst (destEvar body)) evk2 evd' in
+ let evd' = Evd.define evk2 (EConstr.Unsafe.to_constr body) evd in
+ let evd' = update_evar_source (fst (destEvar evd body)) evk2 evd' in
check_evar_instance evd' evk2 body g
with EvarSolvedOnTheFly (evd,c) ->
f env evd pbty ev2 c
@@ -1206,7 +1276,7 @@ let preferred_orientation evd evk1 evk2 =
| _ -> true
let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) =
- let aliases = make_alias_map env in
+ let aliases = make_alias_map env evd in
if preferred_orientation evd evk1 evk2 then
try solve_evar_evar_l2r force f g env evd aliases (opp_problem pbty) ev2 ev1
with CannotProject (evd,ev2) ->
@@ -1223,15 +1293,18 @@ let solve_evar_evar_aux force f g env evd pbty (evk1,args1 as ev1) (evk2,args2 a
let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,args2 as ev2) =
let pbty = if force then None else pbty in
let evi = Evd.find evd evk1 in
+ let downcast evk t evd = downcast evk (EConstr.Unsafe.to_constr t) evd in
let evd =
try
(* ?X : Π Δ. Type i = ?Y : Π Δ'. Type j.
The body of ?X and ?Y just has to be of type Π Δ. Type k for some k <= i, j. *)
let evienv = Evd.evar_env evi in
let ctx1, i = Reduction.dest_arity evienv evi.evar_concl in
+ let ctx1 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx1 in
let evi2 = Evd.find evd evk2 in
let evi2env = Evd.evar_env evi2 in
let ctx2, j = Reduction.dest_arity evi2env evi2.evar_concl in
+ let ctx2 = List.map (fun c -> map_rel_decl EConstr.of_constr c) ctx2 in
let ui, uj = univ_of_sort i, univ_of_sort j in
if i == j || Evd.check_eq evd ui uj
then (* Shortcut, i = j *)
@@ -1253,10 +1326,10 @@ let solve_evar_evar ?(force=false) f g env evd pbty (evk1,args1 as ev1) (evk2,ar
solve_evar_evar_aux force f g env evd pbty ev1 ev2
type conv_fun =
- env -> evar_map -> conv_pb -> constr -> constr -> unification_result
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> unification_result
type conv_fun_bool =
- env -> evar_map -> conv_pb -> constr -> constr -> bool
+ env -> evar_map -> conv_pb -> EConstr.constr -> EConstr.constr -> bool
(* Solve pbs ?e[t1..tn] = ?e[u1..un] which arise often in fixpoint
* definitions. We try to unify the ti with the ui pairwise. The pairs
@@ -1265,7 +1338,13 @@ type conv_fun_bool =
let solve_refl ?(can_drop=false) conv_algo env evd pbty evk argsv1 argsv2 =
let evdref = ref evd in
- if Array.equal (e_eq_constr_univs evdref) argsv1 argsv2 then !evdref else
+ let eq_constr c1 c2 = match EConstr.eq_constr_universes !evdref c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !evdref cstr); true
+ with UniversesDiffer -> false
+ in
+ if Array.equal eq_constr argsv1 argsv2 then !evdref else
(* Filter and restrict if needed *)
let args = Array.map2 (fun a1 a2 -> (a1, a2)) argsv1 argsv2 in
let untypedfilter =
@@ -1297,14 +1376,14 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs =
| Some l ->
let l' =
List.map_filter
- (filter_compatible_candidates conv_algo env evd evi argsv rhs) l in
+ (fun c -> filter_compatible_candidates conv_algo env evd evi argsv rhs (EConstr.of_constr c)) l in
match l' with
| [] -> raise IncompatibleCandidates
| [c,evd] ->
(* solve_candidates might have been called recursively in the mean *)
(* time and the evar been solved by the filtering process *)
if Evd.is_undefined evd evk then
- let evd' = Evd.define evk c evd in
+ let evd' = Evd.define evk (EConstr.Unsafe.to_constr c) evd in
check_evar_instance evd' evk c conv_algo
else evd
| l when List.length l < List.length l' ->
@@ -1313,7 +1392,9 @@ let solve_candidates conv_algo env evd (evk,argsv) rhs =
| l -> evd
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
| Evar (sp,_) when Evar.equal sp n -> raise Occur
| Evar (sp,args as e) ->
@@ -1322,7 +1403,7 @@ let occur_evar_upto_types sigma n c =
else (
seen := Evar.Set.add sp !seen;
Option.iter occur_rec (existential_opt_value sigma e);
- occur_rec (existential_type sigma e))
+ occur_rec (Evd.existential_type sigma e))
| _ -> Constr.iter occur_rec c
in
try occur_rec c; false with Occur -> true
@@ -1350,14 +1431,14 @@ let occur_evar_upto_types sigma n c =
* Σ; Γ, y1..yk |- φ(u1..un) = c /\ x1..xn |- φ(x1..xn)
*)
-exception NotInvertibleUsingOurAlgorithm of constr
+exception NotInvertibleUsingOurAlgorithm of EConstr.constr
exception NotEnoughInformationToProgress of (Id.t * evar_projection) list
-exception NotEnoughInformationEvarEvar of constr
-exception OccurCheckIn of evar_map * constr
+exception NotEnoughInformationEvarEvar of EConstr.constr
+exception OccurCheckIn of evar_map * EConstr.constr
exception MetaOccurInBodyInternal
let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
- let aliases = make_alias_map env in
+ let aliases = make_alias_map env evd in
let evdref = ref evd in
let progress = ref false in
let evi = Evd.find evd evk in
@@ -1374,12 +1455,12 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
| (id,p)::_::_ ->
if choose then (mkVar id, p) else raise (NotUniqueInType sols)
in
- let ty = lazy (Retyping.get_type_of env !evdref t) in
+ let ty = lazy (Retyping.get_type_of env !evdref (of_alias t)) in
let evd = do_projection_effects (evar_define conv_algo ~choose) env ty !evdref p in
evdref := evd;
c
with
- | Not_found -> raise (NotInvertibleUsingOurAlgorithm t)
+ | Not_found -> raise (NotInvertibleUsingOurAlgorithm (of_alias t))
| NotUniqueInType sols ->
if not !progress then
raise (NotEnoughInformationToProgress sols);
@@ -1389,15 +1470,15 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let ty' = instantiate_evar_array evi ty argsv in
let (evd,evar,(evk',argsv' as ev')) =
materialize_evar (evar_define conv_algo ~choose) env !evdref 0 ev ty' in
- let ts = expansions_of_var aliases t in
- let test c = isEvar c || List.mem_f Constr.equal c ts in
+ let ts = expansions_of_var evd aliases t in
+ let test c = isEvar evd c || List.exists (is_alias evd c) ts in
let filter = restrict_upon_filter evd evk test argsv' in
let filter = closure_of_filter evd evk' filter in
let candidates = extract_candidates sols in
let evd = match candidates with
| NoUpdate ->
let evd, ev'' = restrict_applied_evar evd ev' filter NoUpdate in
- Evd.add_conv_pb (Reduction.CONV,env,mkEvar ev'',t) evd
+ add_conv_oriented_pb ~tail:false (None,env,mkEvar ev'',of_alias t) evd
| UpdateWith _ ->
restrict_evar evd evk' filter candidates
in
@@ -1405,21 +1486,20 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
evar in
let rec imitate (env',k as envk) t =
- let t = whd_evar !evdref t in
- match kind_of_term t with
+ match EConstr.kind !evdref t with
| Rel i when i>k ->
let open Context.Rel.Declaration in
(match Environ.lookup_rel (i-k) env' with
- | LocalAssum _ -> project_variable (mkRel (i-k))
+ | LocalAssum _ -> project_variable (RelAlias (i-k))
| LocalDef (_,b,_) ->
- try project_variable (mkRel (i-k))
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i b))
+ try project_variable (RelAlias (i-k))
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk (lift i (EConstr.of_constr b)))
| Var id ->
(match Environ.lookup_named id env' with
- | LocalAssum _ -> project_variable t
+ | LocalAssum _ -> project_variable (VarAlias id)
| LocalDef (_,b,_) ->
- try project_variable t
- with NotInvertibleUsingOurAlgorithm _ -> imitate envk b)
+ try project_variable (VarAlias id)
+ with NotInvertibleUsingOurAlgorithm _ -> imitate envk (EConstr.of_constr b))
| LetIn (na,b,u,c) ->
imitate envk (subst1 b c)
| Evar (evk',args' as ev') ->
@@ -1446,7 +1526,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
(* Try to project (a restriction of) the left evar ... *)
try
let evd,body = project_evar_on_evar false conv_algo env' evd aliases 0 None ev'' ev' in
- let evd = Evd.define evk' body evd in
+ let evd = Evd.define evk' (EConstr.Unsafe.to_constr body) evd in
check_evar_instance evd evk' body conv_algo
with
| EvarSolvedOnTheFly _ -> assert false (* ev has no candidates *)
@@ -1458,9 +1538,9 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
| _ ->
progress := true;
match
- let c,args = decompose_app_vect t in
- match kind_of_term c with
- | Construct (cstr,u) when noccur_between 1 k t ->
+ let c,args = decompose_app_vect !evdref t in
+ match EConstr.kind !evdref c with
+ | Construct (cstr,u) when noccur_between !evdref 1 k t ->
(* This is common case when inferring the return clause of match *)
(* (currently rudimentary: we do not treat the case of multiple *)
(* possible inversions; we do not treat overlap with a possible *)
@@ -1475,7 +1555,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
let candidates =
try
let t =
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1)
imitate envk t in
t::l
with e when CErrors.noncritical e -> l in
@@ -1488,7 +1568,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
evar'')
| None ->
(* Evar/Rigid problem (or assimilated if not normal): we "imitate" *)
- map_constr_with_full_binders (fun d (env,k) -> push_rel d env, k+1)
+ map_constr_with_full_binders !evdref (fun d (env,k) -> push_rel d env, k+1)
imitate envk t
in
let rhs = whd_beta evd rhs (* heuristic *) in
@@ -1500,16 +1580,16 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
| (decl :: ctxt'), (c :: s') ->
let id = get_id decl in
names := Idset.add id !names;
- isVarId id c && is_id_subst ctxt' s'
+ isVarId evd id c && is_id_subst ctxt' s'
| [], [] -> true
| _ -> false
in
is_id_subst filter_ctxt (Array.to_list argsv) &&
- closed0 rhs &&
- Idset.subset (collect_vars rhs) !names
+ closed0 evd rhs &&
+ Idset.subset (collect_vars evd rhs) !names
in
let body =
- if fast rhs then nf_evar evd rhs
+ if fast rhs then EConstr.of_constr (EConstr.to_constr evd rhs) (** FIXME? *)
else
let t' = imitate (env,0) rhs in
if !progress then
@@ -1525,7 +1605,7 @@ let rec invert_definition conv_algo choose env evd pbty (evk,argsv as ev) rhs =
*)
and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
- match kind_of_term rhs with
+ match EConstr.kind evd rhs with
| Evar (evk2,argsv2 as ev2) ->
if Evar.equal evk evk2 then
solve_refl ~can_drop:choose
@@ -1538,7 +1618,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
with NoCandidates ->
try
let (evd',body) = invert_definition conv_algo choose env evd pbty ev rhs in
- if occur_meta body then raise MetaOccurInBodyInternal;
+ if occur_meta evd' body then raise MetaOccurInBodyInternal;
(* invert_definition may have instantiate some evars of rhs with evk *)
(* so we recheck acyclicity *)
if occur_evar_upto_types evd' evk body then raise (OccurCheckIn (evd',body));
@@ -1561,7 +1641,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
print_constr body);
raise e in*)
let evd' = check_evar_instance evd' evk body conv_algo in
- Evd.define evk body evd'
+ Evd.define evk (EConstr.Unsafe.to_constr body) evd'
with
| NotEnoughInformationToProgress sols ->
postpone_non_unique_projection env evd pbty ev sols rhs
@@ -1574,7 +1654,7 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
| OccurCheckIn (evd,rhs) ->
(* last chance: rhs actually reduces to ev *)
let c = whd_all env evd rhs in
- match kind_of_term c with
+ match EConstr.kind evd c with
| Evar (evk',argsv2) when Evar.equal evk evk' ->
solve_refl (fun env sigma pb c c' -> is_fconv pb env sigma c c')
env evd pbty evk argsv argsv2
@@ -1607,17 +1687,19 @@ and evar_define conv_algo ?(choose=false) env evd pbty (evk,argsv as ev) rhs =
* ass.
*)
-let status_changed lev (pbty,_,t1,t2) =
- (try Evar.Set.mem (head_evar t1) lev with NoHeadEvar -> false) ||
- (try Evar.Set.mem (head_evar t2) lev with NoHeadEvar -> false)
+let status_changed evd lev (pbty,_,t1,t2) =
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
+ (try Evar.Set.mem (head_evar evd t1) lev with NoHeadEvar -> false) ||
+ (try Evar.Set.mem (head_evar evd t2) lev with NoHeadEvar -> false)
let reconsider_unif_constraints conv_algo evd =
- let (evd,pbs) = extract_changed_conv_pbs evd status_changed in
+ let (evd,pbs) = extract_changed_conv_pbs evd (status_changed evd) in
List.fold_left
(fun p (pbty,env,t1,t2 as x) ->
match p with
| Success evd ->
- (match conv_algo env evd pbty t1 t2 with
+ (match conv_algo env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2) with
| Success _ as x -> x
| UnifFailure (i,e) -> UnifFailure (i,CannotSolveConstraint (x,e)))
| UnifFailure _ as x -> x)
diff --git a/pretyping/evarsolve.mli b/pretyping/evarsolve.mli
index b6bdc0788..02b5c59d2 100644
--- a/pretyping/evarsolve.mli
+++ b/pretyping/evarsolve.mli
@@ -7,9 +7,14 @@
(************************************************************************)
open Term
+open EConstr
open Evd
open Environ
+type alias
+
+val of_alias : alias -> EConstr.t
+
type unification_result =
| Success of evar_map
| UnifFailure of evar_map * Pretype_errors.unification_error
@@ -18,7 +23,7 @@ val is_success : unification_result -> bool
(** Replace the vars and rels that are aliases to other vars and rels by
their representative that is most ancient in the context *)
-val expand_vars_in_term : env -> constr -> constr
+val expand_vars_in_term : env -> evar_map -> constr -> constr
(** [evar_define choose env ev c] try to instantiate [ev] with [c] (typed in [env]),
possibly solving related unification problems, possibly leaving open
@@ -60,12 +65,12 @@ val reconsider_conv_pbs : conv_fun -> evar_map -> unification_result
(** @deprecated Alias for [reconsider_unif_constraints] *)
val is_unification_pattern_evar : env -> evar_map -> existential -> constr list ->
- constr -> constr list option
+ constr -> alias list option
val is_unification_pattern : env * int -> evar_map -> constr -> constr list ->
- constr -> constr list option
+ constr -> alias list option
-val solve_pattern_eqn : env -> constr list -> constr -> constr
+val solve_pattern_eqn : env -> evar_map -> alias list -> constr -> constr
val noccur_evar : env -> evar_map -> Evar.t -> constr -> bool
@@ -76,7 +81,7 @@ val check_evar_instance :
evar_map -> existential_key -> constr -> conv_fun -> evar_map
val remove_instance_local_defs :
- evar_map -> existential_key -> constr array -> constr list
+ evar_map -> existential_key -> 'a array -> 'a list
val get_type_of_refresh :
?polyprop:bool -> ?lax:bool -> env -> evar_map -> constr -> evar_map * types
diff --git a/pretyping/find_subterm.ml b/pretyping/find_subterm.ml
index 4b9cf415f..3fc569fc4 100644
--- a/pretyping/find_subterm.ml
+++ b/pretyping/find_subterm.ml
@@ -11,7 +11,7 @@ open Util
open CErrors
open Names
open Locus
-open Term
+open EConstr
open Nameops
open Termops
open Pretype_errors
@@ -63,7 +63,10 @@ let proceed_with_occurrences f occs x =
let map_named_declaration_with_hyploc f hyploc acc decl =
let open Context.Named.Declaration in
- let f = f (Some (NamedDecl.get_id decl, hyploc)) in
+ let f acc typ =
+ let acc, typ = f (Some (NamedDecl.get_id decl, hyploc)) acc typ in
+ acc, typ
+ in
match decl,hyploc with
| LocalAssum (id,_), InHypValueOnly ->
error_occurrences_error (IncorrectInValueOccurrence id)
@@ -82,10 +85,10 @@ let map_named_declaration_with_hyploc f hyploc acc decl =
exception SubtermUnificationError of subterm_unification_error
-exception NotUnifiable of (constr * constr * unification_error) option
+exception NotUnifiable of (EConstr.t * EConstr.t * unification_error) option
type 'a testing_function = {
- match_fun : 'a -> constr -> 'a;
+ match_fun : 'a -> EConstr.constr -> 'a;
merge_fun : 'a -> 'a -> 'a;
mutable testing_state : 'a;
mutable last_found : position_reporting option
@@ -96,7 +99,7 @@ type 'a testing_function = {
(b,l), b=true means no occurrence except the ones in l and b=false,
means all occurrences except the ones in l *)
-let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t =
+let replace_term_occ_gen_modulo sigma occs like_first test bywhat cl occ t =
let (nowhere_except_in,locs) = Locusops.convert_occs occs in
let maxocc = List.fold_right max locs 0 in
let pos = ref occ in
@@ -130,23 +133,23 @@ let replace_term_occ_gen_modulo occs like_first test bywhat cl occ t =
with NotUnifiable _ ->
subst_below k t
and subst_below k t =
- map_constr_with_binders_left_to_right (fun d k -> k+1) substrec k t
+ map_constr_with_binders_left_to_right sigma (fun d k -> k+1) substrec k t
in
let t' = substrec 0 t in
(!pos, t')
-let replace_term_occ_modulo occs test bywhat t =
+let replace_term_occ_modulo evd occs test bywhat t =
let occs',like_first =
match occs with AtOccs occs -> occs,false | LikeFirst -> AllOccurrences,true in
proceed_with_occurrences
- (replace_term_occ_gen_modulo occs' like_first test bywhat None) occs' t
+ (replace_term_occ_gen_modulo evd occs' like_first test bywhat None) occs' t
-let replace_term_occ_decl_modulo occs test bywhat d =
+let replace_term_occ_decl_modulo evd occs test bywhat d =
let (plocs,hyploc),like_first =
match occs with AtOccs occs -> occs,false | LikeFirst -> (AllOccurrences,InHyp),true in
proceed_with_occurrences
(map_named_declaration_with_hyploc
- (replace_term_occ_gen_modulo plocs like_first test bywhat)
+ (replace_term_occ_gen_modulo evd plocs like_first test bywhat)
hyploc)
plocs d
@@ -154,11 +157,12 @@ let replace_term_occ_decl_modulo occs test bywhat d =
let make_eq_univs_test env evd c =
{ match_fun = (fun evd c' ->
- let b, cst = Universes.eq_constr_universes_proj env c c' in
- if b then
+ match EConstr.eq_constr_universes_proj env evd c c' with
+ | None -> raise (NotUnifiable None)
+ | Some cst ->
try Evd.add_universe_constraints evd cst
with Evd.UniversesDiffer -> raise (NotUnifiable None)
- else raise (NotUnifiable None));
+ );
merge_fun = (fun evd _ -> evd);
testing_state = evd;
last_found = None
@@ -167,7 +171,7 @@ let make_eq_univs_test env evd c =
let subst_closed_term_occ env evd occs c t =
let test = make_eq_univs_test env evd c in
let bywhat () = mkRel 1 in
- let t' = replace_term_occ_modulo occs test bywhat t in
+ let t' = replace_term_occ_modulo evd occs test bywhat t in
t', test.testing_state
let subst_closed_term_occ_decl env evd occs c d =
@@ -177,6 +181,6 @@ let subst_closed_term_occ_decl env evd occs c d =
let bywhat () = mkRel 1 in
proceed_with_occurrences
(map_named_declaration_with_hyploc
- (fun _ -> replace_term_occ_gen_modulo plocs like_first test bywhat None)
+ (fun _ -> replace_term_occ_gen_modulo evd plocs like_first test bywhat None)
hyploc) plocs d,
test.testing_state
diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli
index c741ab048..e3d3b74f1 100644
--- a/pretyping/find_subterm.mli
+++ b/pretyping/find_subterm.mli
@@ -11,6 +11,7 @@ open Term
open Evd
open Pretype_errors
open Environ
+open EConstr
(** Finding subterms, possibly up to some unification function,
possibly at some given occurrences *)
@@ -41,15 +42,16 @@ val make_eq_univs_test : env -> evar_map -> constr -> evar_map testing_function
matching subterms at the indicated occurrences [occl] with [mk
()]; it turns a NotUnifiable exception raised by the testing
function into a SubtermUnificationError. *)
-val replace_term_occ_modulo : occurrences or_like_first ->
+val replace_term_occ_modulo : evar_map -> occurrences or_like_first ->
'a testing_function -> (unit -> constr) -> constr -> constr
(** [replace_term_occ_decl_modulo] is similar to
[replace_term_occ_modulo] but for a named_declaration. *)
val replace_term_occ_decl_modulo :
+ evar_map ->
(occurrences * hyp_location_flag) or_like_first ->
'a testing_function -> (unit -> constr) ->
- Context.Named.Declaration.t -> Context.Named.Declaration.t
+ named_declaration -> named_declaration
(** [subst_closed_term_occ occl c d] replaces occurrences of
closed [c] at positions [occl] by [Rel 1] in [d] (see also Note OCC),
@@ -61,7 +63,7 @@ val subst_closed_term_occ : env -> evar_map -> occurrences or_like_first ->
closed [c] at positions [occl] by [Rel 1] in [decl]. *)
val subst_closed_term_occ_decl : env -> evar_map ->
(occurrences * hyp_location_flag) or_like_first ->
- constr -> Context.Named.Declaration.t -> Context.Named.Declaration.t * evar_map
+ constr -> named_declaration -> named_declaration * evar_map
(** Miscellaneous *)
val error_invalid_occurrence : int list -> 'a
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 9cf91a947..c4a74d990 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -40,6 +40,18 @@ type recursion_scheme_error =
exception RecursionSchemeError of recursion_scheme_error
+let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na
+let name_assumption env = function
+| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
+| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+
+let mkLambda_or_LetIn_name env d b = mkLambda_or_LetIn (name_assumption env d) b
+let mkProd_or_LetIn_name env d b = mkProd_or_LetIn (name_assumption env d) b
+let mkLambda_name env (n,a,b) = mkLambda_or_LetIn_name env (LocalAssum (n,a)) b
+let mkProd_name env (n,a,b) = mkProd_or_LetIn_name env (LocalAssum (n,a)) b
+let it_mkProd_or_LetIn_name env b l = List.fold_left (fun c d -> mkProd_or_LetIn_name env d c) b l
+let it_mkLambda_or_LetIn_name env b l = List.fold_left (fun c d -> mkLambda_or_LetIn_name env d c) b l
+
let make_prod_dep dep env = if dep then mkProd_name env else mkProd
let mkLambda_string s t c = mkLambda (Name (Id.of_string s), t, c)
@@ -63,7 +75,7 @@ let check_privacy_block mib =
let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let lnamespar = Vars.subst_instance_context u mib.mind_params_ctxt in
- let indf = make_ind_family(pind, Context.Rel.to_extended_list 0 lnamespar) in
+ let indf = make_ind_family(pind, Context.Rel.to_extended_list mkRel 0 lnamespar) in
let constrs = get_constructors env indf in
let projs = get_projections env indf in
@@ -93,8 +105,8 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
let pbody =
appvect
(mkRel (ndepar + nbprod),
- if dep then Context.Rel.to_extended_vect 0 deparsign
- else Context.Rel.to_extended_vect 1 arsign) in
+ if dep then Context.Rel.to_extended_vect mkRel 0 deparsign
+ else Context.Rel.to_extended_vect mkRel 1 arsign) in
let p =
it_mkLambda_or_LetIn_name env'
((if dep then mkLambda_name env' else mkLambda)
@@ -118,12 +130,13 @@ let mis_make_case_com dep env sigma (ind, u as pind) (mib,mip as specif) kind =
it_mkLambda_or_LetIn_name env' obj deparsign
else
let cs = lift_constructor (k+1) constrs.(k) in
- let t = build_branch_type env dep (mkRel (k+1)) cs in
+ let t = build_branch_type env (Sigma.to_evar_map sigma) dep (mkRel (k+1)) cs in
mkLambda_string "f" t
(add_branch (push_rel (LocalAssum (Anonymous, t)) env) (k+1))
in
let Sigma (s, sigma, p) = Sigma.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env sigma kind in
- let typP = make_arity env' dep indf s in
+ let typP = make_arity env' (Sigma.to_evar_map sigma) dep indf s in
+ let typP = EConstr.Unsafe.to_constr typP in
let c =
it_mkLambda_or_LetIn_name env
(mkLambda_string "P" typP
@@ -153,7 +166,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let nparams = List.length vargs in
let process_pos env depK pk =
let rec prec env i sign p =
- let p',largs = whd_allnolet_stack env sigma p in
+ 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
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
@@ -166,11 +181,12 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let base = applist (lift i pk,realargs) in
if depK then
Reduction.beta_appvect
- base [|applist (mkRel (i+1), Context.Rel.to_extended_list 0 sign)|]
+ base [|applist (mkRel (i+1), Context.Rel.to_extended_list mkRel 0 sign)|]
else
base
| _ ->
- let t' = whd_all env sigma p in
+ 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
else prec env i sign t'
in
@@ -229,7 +245,9 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let process_pos env fk =
let rec prec env i hyps p =
- let p',largs = whd_allnolet_stack env sigma p in
+ 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
| Prod (n,t,c) ->
let d = LocalAssum (n,t) in
@@ -239,10 +257,11 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
let realargs = List.skipn nparrec largs
- and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect 0 hyps) in
+ and arg = appvect (mkRel (i+1), Context.Rel.to_extended_vect mkRel 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ ->
- let t' = whd_all env sigma p in
+ 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
else prec env i hyps t'
in
@@ -261,7 +280,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
| None ->
mkLambda_name env
(n,t,process_constr (push_rel d env) (i+1)
- (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1)])))
+ (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1)])))))
(cprest,rest))
| Some(_,f_0) ->
let nF = lift (i+1+decF) f_0 in
@@ -269,7 +288,7 @@ let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let arg = process_pos env' nF (lift 1 t) in
mkLambda_name env
(n,t,process_constr env' (i+1)
- (whd_beta Evd.empty (applist (lift 1 f, [(mkRel 1); arg])))
+ (EConstr.Unsafe.to_constr (whd_beta Evd.empty (EConstr.of_constr (applist (lift 1 f, [(mkRel 1); arg])))))
(cprest,rest)))
| (LocalDef (n,c,t) as d)::cprest, rest ->
mkLetIn
@@ -317,7 +336,7 @@ let mis_make_indrec env sigma listdepkind mib u =
(* arity in the context of the fixpoint, i.e.
P1..P_nrec f1..f_nbconstruct *)
- let args = Context.Rel.to_extended_list (nrec+nbconstruct) lnamesparrec in
+ let args = Context.Rel.to_extended_list mkRel (nrec+nbconstruct) lnamesparrec in
let indf = make_ind_family((indi,u),args) in
let arsign,_ = get_arity env indf in
@@ -331,15 +350,15 @@ let mis_make_indrec env sigma listdepkind mib u =
(* constructors in context of the Cases expr, i.e.
P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
- let args' = Context.Rel.to_extended_list (dect+nrec) lnamesparrec in
- let args'' = Context.Rel.to_extended_list ndepar lnonparrec in
+ let args' = Context.Rel.to_extended_list mkRel (dect+nrec) lnamesparrec in
+ let args'' = Context.Rel.to_extended_list mkRel ndepar lnonparrec in
let indf' = make_ind_family((indi,u),args'@args'') in
let branches =
let constrs = get_constructors env indf' in
let fi = Termops.rel_vect (dect-i-nctyi) nctyi in
let vecfi = Array.map
- (fun f -> appvect (f, Context.Rel.to_extended_vect ndepar lnonparrec))
+ (fun f -> appvect (f, Context.Rel.to_extended_vect mkRel ndepar lnonparrec))
fi
in
Array.map3
@@ -360,9 +379,9 @@ let mis_make_indrec env sigma listdepkind mib u =
let deparsign' = LocalAssum (Anonymous,depind')::arsign' in
let pargs =
- let nrpar = Context.Rel.to_extended_list (2*ndepar) lnonparrec
- and nrar = if dep then Context.Rel.to_extended_list 0 deparsign'
- else Context.Rel.to_extended_list 1 arsign'
+ let nrpar = Context.Rel.to_extended_list mkRel (2*ndepar) lnonparrec
+ and nrar = if dep then Context.Rel.to_extended_list mkRel 0 deparsign'
+ else Context.Rel.to_extended_list mkRel 1 arsign'
in nrpar@nrar
in
@@ -378,9 +397,10 @@ let mis_make_indrec env sigma listdepkind mib u =
arsign'
in
let obj =
- Inductiveops.make_case_or_project env indf ci pred
- (mkRel 1) branches
+ Inductiveops.make_case_or_project env !evdref indf ci (EConstr.of_constr pred)
+ (EConstr.mkRel 1) (Array.map EConstr.of_constr branches)
in
+ let obj = EConstr.to_constr !evdref obj in
it_mkLambda_or_LetIn_name env obj
(Termops.lift_rel_context nrec deparsign)
in
@@ -389,8 +409,8 @@ let mis_make_indrec env sigma listdepkind mib u =
let typtyi =
let concl =
- let pargs = if dep then Context.Rel.to_extended_vect 0 deparsign
- else Context.Rel.to_extended_vect 1 arsign
+ let pargs = if dep then Context.Rel.to_extended_vect mkRel 0 deparsign
+ else Context.Rel.to_extended_vect mkRel 1 arsign
in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs)
in it_mkProd_or_LetIn_name env
concl
@@ -417,7 +437,7 @@ let mis_make_indrec env sigma listdepkind mib u =
else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
let recarg = recargpar@recarg in
- let vargs = Context.Rel.to_extended_list (nrec+i+j) lnamesparrec in
+ let vargs = Context.Rel.to_extended_list mkRel (nrec+i+j) lnamesparrec in
let cs = get_constructor ((indi,u),mibi,mipi,vargs) (j+1) in
let p_0 =
type_rec_branch
@@ -431,12 +451,13 @@ let mis_make_indrec env sigma listdepkind mib u =
in
let rec put_arity env i = function
| ((indi,u),_,_,dep,kinds)::rest ->
- let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list i lnamesparrec) in
+ let indf = make_ind_family ((indi,u), Context.Rel.to_extended_list mkRel i lnamesparrec) in
let s =
Evarutil.evd_comb1 (Evd.fresh_sort_in_family ~rigid:Evd.univ_flexible_alg env)
evdref kinds
in
- let typP = make_arity env dep indf s in
+ let typP = make_arity env !evdref dep indf s in
+ let typP = EConstr.Unsafe.to_constr typP in
mkLambda_string "P" typP
(put_arity (push_rel (LocalAssum (Anonymous,typP)) env) (i+1) rest)
| [] ->
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index ac6d775e3..5b42add28 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -58,21 +58,24 @@ let lift_inductive_family n = liftn_inductive_family n 1
let substnl_ind_family l n = map_ind_family (substnl l n)
-type inductive_type = IndType of inductive_family * constr list
+type inductive_type = IndType of inductive_family * EConstr.constr list
let make_ind_type (indf, realargs) = IndType (indf,realargs)
let dest_ind_type (IndType (indf,realargs)) = (indf,realargs)
let map_inductive_type f (IndType (indf, realargs)) =
- IndType (map_ind_family f indf, List.map f realargs)
+ let f' c = EConstr.Unsafe.to_constr (f (EConstr.of_constr c)) in
+ IndType (map_ind_family f' indf, List.map f realargs)
-let liftn_inductive_type n d = map_inductive_type (liftn n d)
+let liftn_inductive_type n d = map_inductive_type (EConstr.Vars.liftn n d)
let lift_inductive_type n = liftn_inductive_type n 1
-let substnl_ind_type l n = map_inductive_type (substnl l n)
+let substnl_ind_type l n = map_inductive_type (EConstr.Vars.substnl l n)
let mkAppliedInd (IndType ((ind,params), realargs)) =
- applist (mkIndU ind,params@realargs)
+ let open EConstr in
+ let ind = on_snd EInstance.make ind in
+ applist (mkIndU ind, (List.map EConstr.of_constr params)@realargs)
(* Does not consider imbricated or mutually recursive types *)
let mis_is_recursive_subset listind rarg =
@@ -343,24 +346,25 @@ let get_projections env (ind,params) =
| Some (Some (id, projs, pbs)) -> Some projs
| _ -> None
-let make_case_or_project env indf ci pred c branches =
+let make_case_or_project env sigma indf ci pred c branches =
+ let open EConstr in
let projs = get_projections env indf in
match projs with
| None -> (mkCase (ci, pred, c, branches))
| Some ps ->
assert(Array.length branches == 1);
let () =
- let _, _, t = destLambda pred in
+ let _, _, t = destLambda sigma pred in
let (ind, _), _ = dest_ind_family indf in
let mib, _ = Inductive.lookup_mind_specif env ind in
- if (* dependent *) not (noccurn 1 t) &&
+ if (* dependent *) not (Vars.noccurn sigma 1 t) &&
not (has_dependent_elim mib) then
user_err ~hdr:"make_case_or_project"
Pp.(str"Dependent case analysis not allowed" ++
str" on inductive type " ++ Names.MutInd.print (fst ind))
in
let branch = branches.(0) in
- let ctx, br = decompose_lam_n_assum (Array.length ps) branch in
+ let ctx, br = decompose_lam_n_assum sigma (Array.length ps) branch in
let n, subst =
List.fold_right
(fun decl (i, subst) ->
@@ -368,9 +372,9 @@ let make_case_or_project env indf ci pred c branches =
| LocalAssum (na, t) ->
let t = mkProj (Projection.make ps.(i) true, c) in
(i + 1, t :: subst)
- | LocalDef (na, b, t) -> (i, substl subst b :: subst))
+ | LocalDef (na, b, t) -> (i, Vars.substl subst b :: subst))
ctx (0, [])
- in substl subst br
+ in Vars.substl subst br
(* substitution in a signature *)
@@ -408,51 +412,56 @@ let build_dependent_constructor cs =
applist
(mkConstructU cs.cs_cstr,
(List.map (lift cs.cs_nargs) cs.cs_params)
- @(Context.Rel.to_extended_list 0 cs.cs_args))
+ @(Context.Rel.to_extended_list mkRel 0 cs.cs_args))
let build_dependent_inductive env ((ind, params) as indf) =
let arsign,_ = get_arity env indf in
let nrealargs = List.length arsign in
applist
(mkIndU ind,
- (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list 0 arsign))
+ (List.map (lift nrealargs) params)@(Context.Rel.to_extended_list mkRel 0 arsign))
(* builds the arity of an elimination predicate in sort [s] *)
-let make_arity_signature env dep indf =
+let make_arity_signature env sigma dep indf =
let (arsign,_) = get_arity env indf in
+ let arsign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) arsign in
if dep then
(* We need names everywhere *)
- Namegen.name_context env
- ((LocalAssum (Anonymous,build_dependent_inductive env indf))::arsign)
+ Namegen.name_context env sigma
+ ((LocalAssum (Anonymous,EConstr.of_constr (build_dependent_inductive env indf)))::arsign)
(* Costly: would be better to name once for all at definition time *)
else
(* No need to enforce names *)
arsign
-let make_arity env dep indf s = mkArity (make_arity_signature env dep indf, s)
+let make_arity env sigma dep indf s =
+ let open EConstr in
+ it_mkProd_or_LetIn (mkSort s) (make_arity_signature env sigma dep indf)
(* [p] is the predicate and [cs] a constructor summary *)
-let build_branch_type env dep p cs =
+let build_branch_type env sigma dep p cs =
let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in
if dep then
- Namegen.it_mkProd_or_LetIn_name env
- (applist (base,[build_dependent_constructor cs]))
- cs.cs_args
+ EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma
+ (EConstr.of_constr (applist (base,[build_dependent_constructor cs])))
+ (List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) cs.cs_args))
else
- it_mkProd_or_LetIn base cs.cs_args
+ Term.it_mkProd_or_LetIn base cs.cs_args
(**************************************************)
-let extract_mrectype t =
- let (t, l) = decompose_app t in
- match kind_of_term t with
+let extract_mrectype sigma t =
+ let open EConstr in
+ let (t, l) = decompose_app sigma t in
+ match EConstr.kind sigma t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
let find_mrectype_vect env sigma c =
- let (t, l) = decompose_appvect (whd_all env sigma c) in
- match kind_of_term t with
+ let open EConstr in
+ let (t, l) = Termops.decompose_app_vect sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind -> (ind, l)
| _ -> raise Not_found
@@ -460,28 +469,35 @@ let find_mrectype env sigma c =
let (ind, v) = find_mrectype_vect env sigma c in (ind, Array.to_list v)
let find_rectype env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
- | Ind (ind,u as indu) ->
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
+ | Ind (ind,u) ->
let (mib,mip) = Inductive.lookup_mind_specif env ind in
if mib.mind_nparams > List.length l then raise Not_found;
+ let l = List.map EConstr.Unsafe.to_constr l in
let (par,rargs) = List.chop mib.mind_nparams l in
- IndType((indu, par),rargs)
+ let indu = (ind, EInstance.kind sigma u) in
+ IndType((indu, par),List.map EConstr.of_constr rargs)
| _ -> raise Not_found
let find_inductive env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind
when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite <> Decl_kinds.CoFinite ->
+ let l = List.map EConstr.Unsafe.to_constr l in
(ind, l)
| _ -> raise Not_found
let find_coinductive env sigma c =
- let (t, l) = decompose_app (whd_all env sigma c) in
- match kind_of_term t with
+ let open EConstr in
+ let (t, l) = decompose_app sigma (whd_all env sigma c) in
+ match EConstr.kind sigma t with
| Ind ind
when (fst (Inductive.lookup_mind_specif env (fst ind))).mind_finite == Decl_kinds.CoFinite ->
+ let l = List.map EConstr.Unsafe.to_constr l in
(ind, l)
| _ -> raise Not_found
@@ -490,12 +506,12 @@ let find_coinductive env sigma c =
(* find appropriate names for pattern variables. Useful in the Case
and Inversion (case_then_using et case_nodep_then_using) tactics. *)
-let is_predicate_explicitly_dep env pred arsign =
+let is_predicate_explicitly_dep env sigma pred arsign =
let rec srec env pval arsign =
- let pv' = whd_all env Evd.empty pval in
- match kind_of_term pv', arsign with
+ let pv' = whd_all env sigma pval in
+ match EConstr.kind sigma pv', arsign with
| Lambda (na,t,b), (LocalAssum _)::arsign ->
- srec (push_rel_assum (na,t) env) b arsign
+ srec (push_rel_assum (na, t) env) b arsign
| Lambda (na,_,t), _ ->
(* The following code has an impact on the introduction names
@@ -525,17 +541,18 @@ let is_predicate_explicitly_dep env pred arsign =
| _ -> anomaly (Pp.str "Non eta-expanded dep-expanded \"match\" predicate")
in
- srec env pred arsign
+ srec env (EConstr.of_constr pred) arsign
-let is_elim_predicate_explicitly_dependent env pred indf =
+let is_elim_predicate_explicitly_dependent env sigma pred indf =
let arsign,_ = get_arity env indf in
- is_predicate_explicitly_dep env pred arsign
+ is_predicate_explicitly_dep env sigma pred arsign
-let set_names env n brty =
- let (ctxt,cl) = decompose_prod_n_assum n brty in
- Namegen.it_mkProd_or_LetIn_name env cl ctxt
+let set_names env sigma n brty =
+ let open EConstr in
+ let (ctxt,cl) = decompose_prod_n_assum sigma n brty in
+ EConstr.Unsafe.to_constr (Namegen.it_mkProd_or_LetIn_name env sigma cl ctxt)
-let set_pattern_names env ind brv =
+let set_pattern_names env sigma ind brv =
let (mib,mip) = Inductive.lookup_mind_specif env ind in
let arities =
Array.map
@@ -543,10 +560,11 @@ let set_pattern_names env ind brv =
Context.Rel.length ((prod_assum c)) -
mib.mind_nparams)
mip.mind_nf_lc in
- Array.map2 (set_names env) arities brv
+ Array.map2 (set_names env sigma) arities brv
-let type_case_branches_with_names env indspec p c =
+let type_case_branches_with_names env sigma indspec p c =
let (ind,args) = indspec in
+ let args = List.map EConstr.Unsafe.to_constr args in
let (mib,mip as specif) = Inductive.lookup_mind_specif env (fst ind) in
let nparams = mib.mind_nparams in
let (params,realargs) = List.chop nparams args in
@@ -554,8 +572,8 @@ let type_case_branches_with_names env indspec p c =
(* Build case type *)
let conclty = lambda_appvect_assum (mip.mind_nrealdecls+1) p (Array.of_list (realargs@[c])) in
(* Adjust names *)
- if is_elim_predicate_explicitly_dependent env p (ind,params) then
- (set_pattern_names env (fst ind) lbrty, conclty)
+ if is_elim_predicate_explicitly_dependent env sigma p (ind,params) then
+ (set_pattern_names env sigma (fst ind) (Array.map EConstr.of_constr lbrty), conclty)
else (lbrty, conclty)
(* Type of Case predicates *)
@@ -563,7 +581,7 @@ let arity_of_case_predicate env (ind,params) dep k =
let arsign,_ = get_arity env (ind,params) in
let mind = build_dependent_inductive env (ind,params) in
let concl = if dep then mkArrow mind (mkSort k) else mkSort k in
- it_mkProd_or_LetIn concl arsign
+ Term.it_mkProd_or_LetIn concl arsign
(***********************************************)
(* Inferring the sort of parameters of a polymorphic inductive type
@@ -598,17 +616,19 @@ let rec instantiate_universes env evdref scl is = function
let type_of_inductive_knowing_conclusion env sigma ((mib,mip),u) conclty =
match mip.mind_arity with
- | RegularArity s -> sigma, subst_instance_constr u s.mind_user_arity
+ | RegularArity s -> sigma, EConstr.of_constr (subst_instance_constr u s.mind_user_arity)
| TemplateArity ar ->
- let _,scl = Reduction.dest_arity env conclty in
+ let _,scl = splay_arity env sigma conclty in
+ let scl = EConstr.ESorts.kind sigma scl in
let ctx = List.rev mip.mind_arity_ctxt in
let evdref = ref sigma in
let ctx =
instantiate_universes
env evdref scl ar.template_level (ctx,ar.template_param_levels) in
- !evdref, mkArity (List.rev ctx,scl)
+ !evdref, EConstr.of_constr (mkArity (List.rev ctx,scl))
let type_of_projection_knowing_arg env sigma p c ty =
+ let c = EConstr.Unsafe.to_constr c in
let IndType(pars,realargs) =
try find_rectype env sigma ty
with Not_found ->
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 7bd616591..bdb6f996b 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -37,15 +37,15 @@ val substnl_ind_family :
constr list -> int -> inductive_family -> inductive_family
(** An inductive type with its parameters and real arguments *)
-type inductive_type = IndType of inductive_family * constr list
-val make_ind_type : inductive_family * constr list -> inductive_type
-val dest_ind_type : inductive_type -> inductive_family * constr list
-val map_inductive_type : (constr -> constr) -> inductive_type -> inductive_type
+type inductive_type = IndType of inductive_family * EConstr.constr list
+val make_ind_type : inductive_family * EConstr.constr list -> inductive_type
+val dest_ind_type : inductive_type -> inductive_family * EConstr.constr list
+val map_inductive_type : (EConstr.constr -> EConstr.constr) -> inductive_type -> inductive_type
val liftn_inductive_type : int -> int -> inductive_type -> inductive_type
val lift_inductive_type : int -> inductive_type -> inductive_type
-val substnl_ind_type : constr list -> int -> inductive_type -> inductive_type
+val substnl_ind_type : EConstr.constr list -> int -> inductive_type -> inductive_type
-val mkAppliedInd : inductive_type -> constr
+val mkAppliedInd : inductive_type -> EConstr.constr
val mis_is_recursive_subset : int list -> wf_paths -> bool
val mis_is_recursive :
inductive * mutual_inductive_body * one_inductive_body -> bool
@@ -130,7 +130,7 @@ val has_dependent_elim : mutual_inductive_body -> bool
val projection_nparams : projection -> int
val projection_nparams_env : env -> projection -> int
val type_of_projection_knowing_arg : env -> evar_map -> Projection.t ->
- constr -> types -> types
+ EConstr.t -> EConstr.types -> types
(** Extract information from an inductive family *)
@@ -156,17 +156,17 @@ 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 -> bool -> inductive_family -> Context.Rel.t
-val make_arity : env -> bool -> inductive_family -> sorts -> types
-val build_branch_type : env -> bool -> constr -> constructor_summary -> types
+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 build_branch_type : env -> evar_map -> bool -> constr -> constructor_summary -> types
(** Raise [Not_found] if not given a valid inductive type *)
-val extract_mrectype : constr -> pinductive * constr list
-val find_mrectype : env -> evar_map -> types -> pinductive * constr list
-val find_mrectype_vect : env -> evar_map -> types -> pinductive * constr array
-val find_rectype : env -> evar_map -> types -> inductive_type
-val find_inductive : env -> evar_map -> types -> pinductive * constr list
-val find_coinductive : env -> evar_map -> types -> pinductive * constr list
+val extract_mrectype : evar_map -> EConstr.t -> (inductive * EConstr.EInstance.t) * EConstr.constr list
+val find_mrectype : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr list
+val find_mrectype_vect : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * EConstr.constr array
+val find_rectype : env -> evar_map -> EConstr.types -> inductive_type
+val find_inductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list
+val find_coinductive : env -> evar_map -> EConstr.types -> (inductive * EConstr.EInstance.t) * constr list
(********************)
@@ -175,7 +175,7 @@ val arity_of_case_predicate :
env -> inductive_family -> bool -> sorts -> types
val type_case_branches_with_names :
- env -> pinductive * constr list -> constr -> constr -> types array * types
+ env -> evar_map -> pinductive * EConstr.constr list -> constr -> constr -> types array * types
(** Annotation for cases *)
val make_case_info : env -> inductive -> case_style -> case_info
@@ -185,8 +185,8 @@ val make_case_info : env -> inductive -> case_style -> case_info
Fail with an error if the elimination is dependent while the
inductive type does not allow dependent elimination. *)
val make_case_or_project :
- env -> inductive_family -> case_info ->
- (* pred *) constr -> (* term *) constr -> (* branches *) constr array -> constr
+ env -> evar_map -> inductive_family -> case_info ->
+ (* pred *) EConstr.constr -> (* term *) EConstr.constr -> (* branches *) EConstr.constr array -> EConstr.constr
(*i Compatibility
val make_default_case_info : env -> case_style -> inductive -> case_info
@@ -195,7 +195,7 @@ i*)
(********************)
val type_of_inductive_knowing_conclusion :
- env -> evar_map -> Inductive.mind_specif puniverses -> types -> evar_map * types
+ env -> evar_map -> Inductive.mind_specif puniverses -> EConstr.types -> evar_map * EConstr.types
(********************)
val control_only_guard : env -> types -> unit
diff --git a/pretyping/miscops.ml b/pretyping/miscops.ml
index 142e430ff..7fe81c9a4 100644
--- a/pretyping/miscops.ml
+++ b/pretyping/miscops.ml
@@ -58,3 +58,16 @@ let map_red_expr_gen f g h = function
| CbvNative occs_o -> CbvNative (Option.map (map_occs (map_union g h)) occs_o)
| Cbn flags -> Cbn (map_flags g flags)
| ExtraRedExpr _ | Red _ | Hnf as x -> x
+
+(** Mapping bindings *)
+
+let map_explicit_bindings f l =
+ let map (loc, hyp, x) = (loc, hyp, f x) in
+ List.map map l
+
+let map_bindings f = function
+| ImplicitBindings l -> ImplicitBindings (List.map f l)
+| ExplicitBindings expl -> ExplicitBindings (map_explicit_bindings f expl)
+| NoBindings -> NoBindings
+
+let map_with_bindings f (x, bl) = (f x, map_bindings f bl)
diff --git a/pretyping/miscops.mli b/pretyping/miscops.mli
index 337473a6f..f30dc1a4b 100644
--- a/pretyping/miscops.mli
+++ b/pretyping/miscops.mli
@@ -27,3 +27,8 @@ val intro_pattern_naming_eq :
val map_red_expr_gen : ('a -> 'd) -> ('b -> 'e) -> ('c -> 'f) ->
('a,'b,'c) red_expr_gen -> ('d,'e,'f) red_expr_gen
+
+(** Mapping bindings *)
+
+val map_bindings : ('a -> 'b) -> 'a bindings -> 'b bindings
+val map_with_bindings : ('a -> 'b) -> 'a with_bindings -> 'b with_bindings
diff --git a/pretyping/nativenorm.ml b/pretyping/nativenorm.ml
index 1e5f12b20..0228f63cd 100644
--- a/pretyping/nativenorm.ml
+++ b/pretyping/nativenorm.ml
@@ -92,13 +92,15 @@ let construct_of_constr_const env tag typ =
let construct_of_constr_block = construct_of_constr false
-let build_branches_type env (mind,_ as _ind) mib mip u params dep p =
+let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p =
let rtbl = mip.mind_reloc_tbl in
(* [build_one_branch i cty] construit le type de la ieme branche (commence
a 0) et les lambda correspondant aux realargs *)
let build_one_branch i cty =
let typi = type_constructor mind mib u cty params in
- let decl,indapp = Reductionops.splay_prod env Evd.empty typi in
+ let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in
+ let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in
+ let indapp = EConstr.Unsafe.to_constr indapp in
let decl_with_letin,_ = decompose_prod_assum typi in
let ind,cargs = find_rectype_a env indapp in
let nparams = Array.length params in
@@ -172,9 +174,9 @@ let branch_of_switch lvl ans bs =
bs ci in
Array.init (Array.length tbl) branch
-let rec nf_val env v typ =
+let rec nf_val env sigma v typ =
match kind_of_value v with
- | Vaccu accu -> nf_accu env accu
+ | Vaccu accu -> nf_accu env sigma accu
| Vfun f ->
let lvl = nb_rel env in
let name,dom,codom =
@@ -184,44 +186,44 @@ let rec nf_val env v typ =
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
let env = push_rel (LocalAssum (name,dom)) env in
- let body = nf_val env (f (mk_rel_accu lvl)) codom in
+ let body = nf_val env sigma (f (mk_rel_accu lvl)) codom in
mkLambda(name,dom,body)
| Vconst n -> construct_of_constr_const env n typ
| Vblock b ->
let capp,ctyp = construct_of_constr_block env (block_tag b) typ in
- let args = nf_bargs env b ctyp in
+ let args = nf_bargs env sigma b ctyp in
mkApp(capp,args)
-and nf_type env v =
+and nf_type env sigma v =
match kind_of_value v with
- | Vaccu accu -> nf_accu env accu
+ | Vaccu accu -> nf_accu env sigma accu
| _ -> assert false
-and nf_type_sort env v =
+and nf_type_sort env sigma v =
match kind_of_value v with
| Vaccu accu ->
- let t,s = nf_accu_type env accu in
+ let t,s = nf_accu_type env sigma accu in
let s = try destSort s with DestKO -> assert false in
t, s
| _ -> assert false
-and nf_accu env accu =
+and nf_accu env sigma accu =
let atom = atom_of_accu accu in
- if Int.equal (accu_nargs accu) 0 then nf_atom env atom
+ if Int.equal (accu_nargs accu) 0 then nf_atom env sigma atom
else
- let a,typ = nf_atom_type env atom in
- let _, args = nf_args env accu typ in
+ let a,typ = nf_atom_type env sigma atom in
+ let _, args = nf_args env sigma accu typ in
mkApp(a,Array.of_list args)
-and nf_accu_type env accu =
+and nf_accu_type env sigma accu =
let atom = atom_of_accu accu in
- if Int.equal (accu_nargs accu) 0 then nf_atom_type env atom
+ if Int.equal (accu_nargs accu) 0 then nf_atom_type env sigma atom
else
- let a,typ = nf_atom_type env atom in
- let t, args = nf_args env accu typ in
+ let a,typ = nf_atom_type env sigma atom in
+ let t, args = nf_args env sigma accu typ in
mkApp(a,Array.of_list args), t
-and nf_args env accu t =
+and nf_args env sigma accu t =
let aux arg (t,l) =
let _,dom,codom =
try decompose_prod env t with
@@ -229,13 +231,13 @@ and nf_args env accu t =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let c = nf_val env arg dom in
+ let c = nf_val env sigma arg dom in
(subst1 c codom, c::l)
in
let t,l = List.fold_right aux (args_of_accu accu) (t,[]) in
t, List.rev l
-and nf_bargs env b t =
+and nf_bargs env sigma b t =
let t = ref t in
let len = block_size b in
Array.init len
@@ -246,10 +248,10 @@ and nf_bargs env b t =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let c = nf_val env (block_field b i) dom in
+ let c = nf_val env sigma (block_field b i) dom in
t := subst1 c codom; c)
-and nf_atom env atom =
+and nf_atom env sigma atom =
match atom with
| Arel i -> mkRel (nb_rel env - i)
| Aconstant cst -> mkConstU cst
@@ -257,19 +259,19 @@ and nf_atom env atom =
| Asort s -> mkSort s
| Avar id -> mkVar id
| Aprod(n,dom,codom) ->
- let dom = nf_type env dom in
+ let dom = nf_type env sigma dom in
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
- let codom = nf_type env (codom vn) in
+ let codom = nf_type env sigma (codom vn) in
mkProd(n,dom,codom)
| Ameta (mv,_) -> mkMeta mv
| Aevar (ev,_) -> mkEvar ev
| Aproj(p,c) ->
- let c = nf_accu env c in
+ let c = nf_accu env sigma c in
mkProj(Projection.make p true,c)
- | _ -> fst (nf_atom_type env atom)
+ | _ -> fst (nf_atom_type env sigma atom)
-and nf_atom_type env atom =
+and nf_atom_type env sigma atom =
match atom with
| Arel i ->
let n = (nb_rel env - i) in
@@ -283,7 +285,7 @@ and nf_atom_type env atom =
| Avar id ->
mkVar id, type_of_var env id
| Acase(ans,accu,p,bs) ->
- let a,ta = nf_accu_type env accu in
+ let a,ta = nf_accu_type env sigma accu in
let ((mind,_),u as ind),allargs = find_rectype_a env ta in
let (mib,mip) = Inductive.lookup_mind_specif env (fst ind) in
let nparams = mib.mind_nparams in
@@ -292,14 +294,14 @@ and nf_atom_type env atom =
hnf_prod_applist env
(Inductiveops.type_of_inductive env ind) (Array.to_list params) in
let pT = whd_all env pT in
- let dep, p = nf_predicate env ind mip params p pT in
+ let dep, p = nf_predicate env sigma ind mip params p pT in
(* Calcul du type des branches *)
- let btypes = build_branches_type env (fst ind) mib mip u params dep p in
+ let btypes = build_branches_type env sigma (fst ind) mib mip u params dep p in
(* calcul des branches *)
let bsw = branch_of_switch (nb_rel env) ans bs in
let mkbranch i v =
let decl,decl_with_letin,codom = btypes.(i) in
- let b = nf_val (Termops.push_rels_assum decl env) v codom in
+ let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in
Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin
in
let branchs = Array.mapi mkbranch bsw in
@@ -307,7 +309,7 @@ and nf_atom_type env atom =
let ci = ans.asw_ci in
mkCase(ci, p, a, branchs), tcase
| Afix(tt,ft,rp,s) ->
- let tt = Array.map (fun t -> nf_type env t) tt in
+ let tt = Array.map (fun t -> nf_type env sigma t) tt in
let name = Array.map (fun _ -> (Name (id_of_string "Ffix"))) tt in
let lvl = nb_rel env in
let nbfix = Array.length ft in
@@ -316,37 +318,37 @@ and nf_atom_type env atom =
let env = push_rec_types (name,tt,[||]) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
- let norm_body i v = nf_val env (napply v fargs) (lift nbfix tt.(i)) in
+ let norm_body i v = nf_val env sigma (napply v fargs) (lift nbfix tt.(i)) in
let ft = Array.mapi norm_body ft in
mkFix((rp,s),(name,tt,ft)), tt.(s)
| Acofix(tt,ft,s,_) | Acofixe(tt,ft,s,_) ->
- let tt = Array.map (nf_type env) tt in
+ let tt = Array.map (nf_type env sigma) tt in
let name = Array.map (fun _ -> (Name (id_of_string "Fcofix"))) tt in
let lvl = nb_rel env in
let fargs = mk_rels_accu lvl (Array.length ft) in
let env = push_rec_types (name,tt,[||]) env in
- let ft = Array.mapi (fun i v -> nf_val env (napply v fargs) tt.(i)) ft in
+ let ft = Array.mapi (fun i v -> nf_val env sigma (napply v fargs) tt.(i)) ft in
mkCoFix(s,(name,tt,ft)), tt.(s)
| Aprod(n,dom,codom) ->
- let dom,s1 = nf_type_sort env dom in
+ let dom,s1 = nf_type_sort env sigma dom in
let vn = mk_rel_accu (nb_rel env) in
let env = push_rel (LocalAssum (n,dom)) env in
- let codom,s2 = nf_type_sort env (codom vn) in
+ let codom,s2 = nf_type_sort env sigma (codom vn) in
mkProd(n,dom,codom), mkSort (sort_of_product env s1 s2)
| Aevar(ev,ty) ->
- let ty = nf_type env ty in
+ let ty = nf_type env sigma ty in
mkEvar ev, ty
| Ameta(mv,ty) ->
- let ty = nf_type env ty in
+ let ty = nf_type env sigma ty in
mkMeta mv, ty
| Aproj(p,c) ->
- let c,tc = nf_accu_type env c in
+ let c,tc = nf_accu_type env sigma c in
let cj = make_judge c tc in
let uj = Typeops.judge_of_projection env (Projection.make p true) cj in
uj.uj_val, uj.uj_type
-and nf_predicate env ind mip params v pT =
+and nf_predicate env sigma ind mip params v pT =
match kind_of_value v, kind_of_term pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
@@ -358,7 +360,7 @@ and nf_predicate env ind mip params v pT =
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
let dep,body =
- nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in
+ nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
| Vfun f, _ ->
let k = nb_rel env in
@@ -368,27 +370,28 @@ and nf_predicate env ind mip params v pT =
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
- let body = nf_type (push_rel (LocalAssum (name,dom)) env) vb in
+ let body = nf_type (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_type env v
+ | _, _ -> false, nf_type env sigma v
let evars_of_evar_map sigma =
{ Nativelambda.evars_val = Evd.existential_opt_value sigma;
Nativelambda.evars_typ = Evd.existential_type sigma;
Nativelambda.evars_metas = Evd.meta_type sigma }
-let native_norm env sigma c ty =
+let native_norm env sigma c ty =
+ let c = EConstr.Unsafe.to_constr c in
+ let ty = EConstr.Unsafe.to_constr ty in
if Coq_config.no_native_compiler then
error "Native_compute reduction has been disabled at configure time."
else
let penv = Environ.pre_env env in
- let sigma = evars_of_evar_map sigma in
(*
Format.eprintf "Numbers of free variables (named): %i\n" (List.length vl1);
Format.eprintf "Numbers of free variables (rel): %i\n" (List.length vl2);
*)
let ml_filename, prefix = Nativelib.get_ml_filename () in
- let code, upd = mk_norm_code penv sigma prefix c in
+ let code, upd = mk_norm_code penv (evars_of_evar_map sigma) prefix c in
match Nativelib.compile ml_filename code with
| true, fn ->
if !Flags.debug then Feedback.msg_debug (Pp.str "Running norm ...");
@@ -397,11 +400,11 @@ let native_norm env sigma c ty =
let t1 = Sys.time () in
let time_info = Format.sprintf "Evaluation done in %.5f@." (t1 -. t0) in
if !Flags.debug then Feedback.msg_debug (Pp.str time_info);
- let res = nf_val env !Nativelib.rt1 ty in
+ let res = nf_val env sigma !Nativelib.rt1 ty in
let t2 = Sys.time () in
let time_info = Format.sprintf "Reification done in %.5f@." (t2 -. t1) in
if !Flags.debug then Feedback.msg_debug (Pp.str time_info);
- res
+ EConstr.of_constr res
| _ -> anomaly (Pp.str "Compilation failure")
let native_conv_generic pb sigma t =
diff --git a/pretyping/nativenorm.mli b/pretyping/nativenorm.mli
index 0b1ce8e51..c899340c8 100644
--- a/pretyping/nativenorm.mli
+++ b/pretyping/nativenorm.mli
@@ -5,7 +5,8 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+
+open EConstr
open Environ
open Evd
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 79765a493..318f94be2 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -116,7 +116,7 @@ let rec head_pattern_bound t =
| PLambda _ -> raise BoundPattern
| PCoFix _ -> anomaly ~label:"head_pattern_bound" (Pp.str "not a type")
-let head_of_constr_reference c = match kind_of_term c with
+let head_of_constr_reference sigma c = match EConstr.kind sigma c with
| Const (sp,_) -> ConstRef sp
| Construct (sp,_) -> ConstructRef sp
| Ind (sp,_) -> IndRef sp
@@ -155,18 +155,14 @@ let pattern_of_constr env sigma t =
| Ind (sp,u) -> PRef (canonical_gr (IndRef sp))
| Construct (sp,u) -> PRef (canonical_gr (ConstructRef sp))
| Proj (p, c) ->
- pattern_of_constr env (Retyping.expand_projection env sigma p c [])
+ pattern_of_constr env (EConstr.Unsafe.to_constr (Retyping.expand_projection env sigma p (EConstr.of_constr c) []))
| Evar (evk,ctxt as ev) ->
(match snd (Evd.evar_source evk sigma) with
| Evar_kinds.MatchingVar (b,id) ->
- let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let () = ignore (pattern_of_constr env ty) in
assert (not b); PMeta (Some id)
| Evar_kinds.GoalEvar ->
PEvar (evk,Array.map (pattern_of_constr env) ctxt)
| _ ->
- let ty = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let () = ignore (pattern_of_constr env ty) in
PMeta None)
| Case (ci,p,a,br) ->
let cip =
@@ -211,6 +207,8 @@ let error_instantiate_pattern id l =
++ strbrk " which " ++ str is ++ strbrk " not bound in the pattern.")
let instantiate_pattern env sigma lvar c =
+ let open EConstr in
+ let open Vars in
let rec aux vars = function
| PVar id as x ->
(try
@@ -222,7 +220,7 @@ let instantiate_pattern env sigma lvar c =
ctx
in
let c = substl inst c in
- pattern_of_constr env sigma c
+ pattern_of_constr env sigma (EConstr.Unsafe.to_constr c)
with Not_found (* List.index failed *) ->
let vars =
List.map_filter (function Name id -> Some id | _ -> None) vars in
diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli
index 1f63565d6..5694d345c 100644
--- a/pretyping/patternops.mli
+++ b/pretyping/patternops.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Globnames
open Glob_term
open Mod_subst
@@ -32,13 +33,13 @@ val head_pattern_bound : constr_pattern -> global_reference
(** [head_of_constr_reference c] assumes [r] denotes a reference and
returns its label; raises an anomaly otherwise *)
-val head_of_constr_reference : Term.constr -> global_reference
+val head_of_constr_reference : Evd.evar_map -> constr -> global_reference
(** [pattern_of_constr c] translates a term [c] with metavariables into
a pattern; currently, no destructor (Cases, Fix, Cofix) and no
existential variable are allowed in [c] *)
-val pattern_of_constr : Environ.env -> Evd.evar_map -> constr -> constr_pattern
+val pattern_of_constr : Environ.env -> Evd.evar_map -> Constr.constr -> constr_pattern
(** [pattern_of_glob_constr l c] translates a term [c] with metavariables into
a pattern; variables bound in [l] are replaced by the pattern to which they
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index 5b0958695..24f6d1689 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -10,6 +10,7 @@ open Util
open Names
open Term
open Environ
+open EConstr
open Type_errors
type unification_error =
@@ -31,6 +32,8 @@ type position_reporting = (position * int) * constr
type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
+type type_error = (constr, types) ptype_error
+
type pretype_error =
(* Old Case *)
| CantFindCaseType of constr
@@ -75,14 +78,19 @@ let error_actual_type ?loc env sigma {uj_val=c;uj_type=actty} expty reason =
raise_pretype_error ?loc
(env, sigma, ActualTypeNotCoercible (j, expty, reason))
+let error_actual_type_core ?loc env sigma {uj_val=c;uj_type=actty} expty =
+ let j = {uj_val=c;uj_type=actty} in
+ raise_type_error ?loc
+ (env, sigma, ActualType (j, expty))
+
let error_cant_apply_not_functional ?loc env sigma rator randl =
raise_type_error ?loc
- (env, sigma, CantApplyNonFunctional (rator, Array.of_list randl))
+ (env, sigma, CantApplyNonFunctional (rator, randl))
let error_cant_apply_bad_type ?loc env sigma (n,c,t) rator randl =
raise_type_error ?loc
(env, sigma,
- CantApplyBadType ((n,c,t), rator, Array.of_list randl))
+ CantApplyBadType ((n,c,t), rator, randl))
let error_ill_formed_branch ?loc env sigma c i actty expty =
raise_type_error
@@ -98,9 +106,16 @@ let error_ill_typed_rec_body ?loc env sigma i na jl tys =
raise_type_error ?loc
(env, sigma, IllTypedRecBody (i, na, jl, tys))
+let error_elim_arity ?loc env sigma pi s c j a =
+ raise_type_error ?loc
+ (env, sigma, ElimArity (pi, s, c, j, a))
+
let error_not_a_type ?loc env sigma j =
raise_type_error ?loc (env, sigma, NotAType j)
+let error_assumption ?loc env sigma j =
+ raise_type_error ?loc (env, sigma, BadAssumption j)
+
(*s Implicit arguments synthesis errors. It is hard to find
a precise location. *)
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index 73f81923f..c303d5d94 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -9,6 +9,7 @@
open Names
open Term
open Environ
+open EConstr
open Type_errors
(** {6 The type of errors raised by the pretyper } *)
@@ -32,6 +33,8 @@ type position_reporting = (position * int) * constr
type subterm_unification_error = bool * position_reporting * position_reporting * (constr * constr * unification_error) option
+type type_error = (constr, types) ptype_error
+
type pretype_error =
(** Old Case *)
| CantFindCaseType of constr
@@ -68,13 +71,16 @@ val error_actual_type :
?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr ->
unification_error -> 'b
+val error_actual_type_core :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> constr -> 'b
+
val error_cant_apply_not_functional :
?loc:Loc.t -> env -> Evd.evar_map ->
- unsafe_judgment -> unsafe_judgment list -> 'b
+ unsafe_judgment -> unsafe_judgment array -> 'b
val error_cant_apply_bad_type :
?loc:Loc.t -> env -> Evd.evar_map -> int * constr * constr ->
- unsafe_judgment -> unsafe_judgment list -> 'b
+ unsafe_judgment -> unsafe_judgment array -> 'b
val error_case_not_inductive :
?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
@@ -91,9 +97,17 @@ val error_ill_typed_rec_body :
?loc:Loc.t -> env -> Evd.evar_map ->
int -> Name.t array -> unsafe_judgment array -> types array -> 'b
+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
+
val error_not_a_type :
?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+val error_assumption :
+ ?loc:Loc.t -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+
val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
(** {6 Implicit arguments synthesis errors } *)
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 27144b279..a042b73c2 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -27,12 +27,14 @@ open Util
open Names
open Evd
open Term
-open Vars
open Termops
-open Reductionops
open Environ
+open EConstr
+open Vars
+open Reductionops
open Type_errors
open Typeops
+open Typing
open Globnames
open Nameops
open Evarutil
@@ -59,7 +61,7 @@ type ltac_var_map = {
ltac_genargs : unbound_ltac_var_map;
}
type glob_constr_ltac_closure = ltac_var_map * glob_constr
-type pure_open_constr = evar_map * constr
+type pure_open_constr = evar_map * EConstr.constr
(************************************************************************)
(* This concerns Cases *)
@@ -77,26 +79,26 @@ type t = {
(** Delay the computation of the evar extended environment *)
}
-let get_extra env =
+let get_extra env sigma =
let open Context.Named.Declaration in
let ids = List.map get_id (named_context env) in
let avoid = List.fold_right Id.Set.add ids Id.Set.empty in
- Context.Rel.fold_outside push_rel_decl_to_named_context
- (Environ.rel_context env) ~init:(empty_csubst, [], avoid, named_context env)
+ Context.Rel.fold_outside (fun d acc -> push_rel_decl_to_named_context sigma d acc)
+ (rel_context env) ~init:(empty_csubst, [], avoid, named_context env)
-let make_env env = { env = env; extra = lazy (get_extra env) }
+let make_env env sigma = { env = env; extra = lazy (get_extra env sigma) }
let rel_context env = rel_context env.env
-let push_rel d env = {
+let push_rel sigma d env = {
env = push_rel d env.env;
- extra = lazy (push_rel_decl_to_named_context d (Lazy.force env.extra));
+ extra = lazy (push_rel_decl_to_named_context sigma d (Lazy.force env.extra));
}
-let pop_rel_context n env = make_env (pop_rel_context n env.env)
+let pop_rel_context n env sigma = make_env (pop_rel_context n env.env) sigma
-let push_rel_context ctx env = {
+let push_rel_context sigma ctx env = {
env = push_rel_context ctx env.env;
- extra = lazy (List.fold_right push_rel_decl_to_named_context ctx (Lazy.force env.extra));
+ extra = lazy (List.fold_right (fun d acc -> push_rel_decl_to_named_context sigma d acc) ctx (Lazy.force env.extra));
}
let lookup_named id env = lookup_named id env.env
@@ -115,9 +117,9 @@ let e_new_evar env evdref ?src ?naming typ =
evdref := Sigma.to_evar_map sigma;
e
-let push_rec_types (lna,typarray,_) env =
+let push_rec_types sigma (lna,typarray,_) env =
let ctxt = Array.map2_i (fun i na t -> Context.Rel.Declaration.LocalAssum (na, lift i t)) lna typarray in
- Array.fold_left (fun e assum -> push_rel assum e) env ctxt
+ Array.fold_left (fun e assum -> push_rel sigma assum e) env ctxt
end
@@ -127,6 +129,10 @@ open ExtraEnv
exception Found of int array
+let nf_fix sigma (nas, cs, ts) =
+ let inj c = EConstr.to_constr sigma c in
+ (nas, Array.map inj cs, Array.map inj ts)
+
let search_guard loc env possible_indexes fixdefs =
(* Standard situation with only one possibility for each fix. *)
(* We treat it separately in order to get proper error msg. *)
@@ -304,7 +310,7 @@ let apply_inference_hook hook evdref frozen = match frozen with
then
try
let sigma, c = hook sigma evk in
- Evd.define evk c sigma
+ Evd.define evk (EConstr.Unsafe.to_constr c) sigma
with Exit ->
sigma
else
@@ -337,17 +343,15 @@ let check_extra_evars_are_solved env current_sigma frozen = match frozen with
let check_evars env initial_sigma sigma c =
let rec proc_rec c =
- match kind_of_term c with
- | Evar (evk,_ as ev) ->
- (match existential_opt_value sigma ev with
- | Some c -> proc_rec c
- | None ->
- if not (Evd.mem initial_sigma evk) then
- let (loc,k) = evar_source evk sigma in
- match k with
- | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
- | _ -> Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None)
- | _ -> Constr.iter proc_rec c
+ match EConstr.kind sigma c with
+ | Evar (evk, _) ->
+ if not (Evd.mem initial_sigma evk) then
+ let (loc,k) = evar_source evk sigma in
+ begin match k with
+ | Evar_kinds.ImplicitArg (gr, (i, id), false) -> ()
+ | _ -> Pretype_errors.error_unsolvable_implicit ~loc env sigma evk None
+ end
+ | _ -> EConstr.iter sigma proc_rec c
in proc_rec c
let check_evars_are_solved env current_sigma frozen =
@@ -411,7 +415,7 @@ let ltac_interp_name { ltac_idents ; ltac_genargs } = function
str"It cannot be used in a binder.")
else n
-let ltac_interp_name_env k0 lvar env =
+let ltac_interp_name_env k0 lvar env sigma =
(* envhd is the initial part of the env when pretype was called first *)
(* (in practice is is probably 0, but we have to grant the
specification of pretype which accepts to start with a non empty
@@ -422,7 +426,7 @@ let ltac_interp_name_env k0 lvar env =
let open Context.Rel.Declaration in
let ctxt' = List.smartmap (map_name (ltac_interp_name lvar)) ctxt in
if List.equal (fun d1 d2 -> Name.equal (get_name d1) (get_name d2)) ctxt ctxt' then env
- else push_rel_context ctxt' (pop_rel_context n env)
+ else push_rel_context sigma ctxt' (pop_rel_context n env sigma)
let invert_ltac_bound_name lvar env id0 id =
let id' = Id.Map.find id lvar.ltac_idents in
@@ -446,7 +450,7 @@ let pretype_id pretype k0 loc env evdref lvar id =
let (n,_,typ) = lookup_rel_id id (rel_context env) in
{ uj_val = mkRel n; uj_type = lift n typ }
with Not_found ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
(* Check if [id] is an ltac variable *)
try
let (ids,c) = Id.Map.find id lvar.ltac_constrs in
@@ -482,9 +486,6 @@ let pretype_id pretype k0 loc env evdref lvar id =
(* [id] not found, standard error message *)
error_var_not_found ~loc id
-let evar_kind_of_term sigma c =
- kind_of_term (whd_evar sigma c)
-
(*************************************************************************)
(* Main pretyping function *)
@@ -516,7 +517,8 @@ let pretype_global loc rigid env evd gr us =
str " universe instances must be greater or equal to Set.");
evd, Some (Univ.Instance.of_array (Array.of_list (List.rev l')))
in
- Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr
+ let (sigma, c) = Evd.fresh_global ~loc ~rigid ?names:instance env.ExtraEnv.env evd gr in
+ (sigma, EConstr.of_constr c)
let pretype_ref loc evdref env ref us =
match ref with
@@ -531,7 +533,7 @@ let pretype_ref loc evdref env ref us =
| ref ->
let evd, c = pretype_global loc univ_flexible env !evdref ref us in
let () = evdref := evd in
- let ty = Typing.unsafe_type_of env.ExtraEnv.env evd c in
+ let ty = unsafe_type_of env.ExtraEnv.env evd c in
make_judge c ty
let judge_of_Type loc evd s =
@@ -591,7 +593,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
inh_conv_coerce_to_tycon loc env evdref j tycon
| GPatVar (loc,(someta,n)) ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
@@ -600,7 +602,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
{ uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty }
| GHole (loc, k, naming, None) ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
@@ -609,7 +611,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
{ uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty }
| GHole (loc, k, _naming, Some arg) ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
| Some ty -> ty
@@ -627,18 +629,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let ty' = pretype_type empty_valcon env evdref lvar ty in
let dcl = LocalAssum (na, ty'.utj_val) in
let dcl' = LocalAssum (ltac_interp_name lvar na,ty'.utj_val) in
- type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl
+ type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl
| (na,bk,Some bd,ty)::bl ->
let ty' = pretype_type empty_valcon env evdref lvar ty in
let bd' = pretype (mk_tycon ty'.utj_val) env evdref lvar bd in
let dcl = LocalDef (na, bd'.uj_val, ty'.utj_val) in
let dcl' = LocalDef (ltac_interp_name lvar na, bd'.uj_val, ty'.utj_val) in
- type_bl (push_rel dcl env) (Context.Rel.add dcl' ctxt) bl in
+ type_bl (push_rel !evdref dcl env) (Context.Rel.add dcl' ctxt) bl in
let ctxtv = Array.map (type_bl env Context.Rel.empty) bl in
let larj =
Array.map2
(fun e ar ->
- pretype_type empty_valcon (push_rel_context e env) evdref lvar ar)
+ pretype_type empty_valcon (push_rel_context !evdref e env) evdref lvar ar)
ctxtv lar in
let lara = Array.map (fun a -> a.utj_val) larj in
let ftys = Array.map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in
@@ -654,23 +656,24 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| None -> true
in
(* Note: bodies are not used by push_rec_types, so [||] is safe *)
- let newenv = push_rec_types (names,ftys,[||]) env in
+ let newenv = push_rec_types !evdref (names,ftys,[||]) env in
let vdefj =
Array.map2_i
(fun i ctxt def ->
(* we lift nbfix times the type in tycon, because of
* the nbfix variables pushed to newenv *)
let (ctxt,ty) =
- decompose_prod_n_assum (Context.Rel.length ctxt)
+ decompose_prod_n_assum !evdref (Context.Rel.length ctxt)
(lift nbfix ftys.(i)) in
- let nenv = push_rel_context ctxt newenv in
+ let nenv = push_rel_context !evdref ctxt newenv in
let j = pretype (mk_tycon ty) nenv evdref lvar def in
{ uj_val = it_mkLambda_or_LetIn j.uj_val ctxt;
uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
ctxtv vdef in
Typing.check_type_fixpoint loc env.ExtraEnv.env evdref names ftys vdefj;
- let ftys = Array.map (nf_evar !evdref) ftys in
- let fdefs = Array.map (fun x -> nf_evar !evdref (j_val x)) vdefj in
+ let nf c = nf_evar !evdref c in
+ let ftys = Array.map nf ftys in (** FIXME *)
+ let fdefs = Array.map (fun x -> nf (j_val x)) vdefj in
let fixj = match fixkind with
| GFix (vn,i) ->
(* First, let's find the guard indexes. *)
@@ -689,12 +692,13 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let fixdecls = (names,ftys,fdefs) in
let indexes =
search_guard
- loc env.ExtraEnv.env possible_indexes fixdecls
+ loc env.ExtraEnv.env possible_indexes (nf_fix !evdref fixdecls)
in
make_judge (mkFix ((indexes,i),fixdecls)) ftys.(i)
| GCoFix i ->
- let cofix = (i,(names,ftys,fdefs)) in
- (try check_cofix env.ExtraEnv.env cofix
+ let fixdecls = (names,ftys,fdefs) in
+ let cofix = (i, fixdecls) in
+ (try check_cofix env.ExtraEnv.env (i, nf_fix !evdref fixdecls)
with reraise ->
let (e, info) = CErrors.push reraise in
let info = Loc.add_loc info loc in
@@ -715,24 +719,24 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
(* Bidirectional typechecking hint:
parameters of a constructor are completely determined
by a typing constraint *)
- if Flags.is_program_mode () && length > 0 && isConstruct fj.uj_val then
+ if Flags.is_program_mode () && length > 0 && isConstruct !evdref fj.uj_val then
match tycon with
| None -> []
| Some ty ->
- let ((ind, i), u) = destConstruct fj.uj_val in
+ let ((ind, i), u) = destConstruct !evdref fj.uj_val in
let npars = inductive_nparams ind in
if Int.equal npars 0 then []
else
try
let IndType (indf, args) = find_rectype env.ExtraEnv.env !evdref ty in
let ((ind',u'),pars) = dest_ind_family indf in
- if eq_ind ind ind' then pars
+ if eq_ind ind ind' then List.map EConstr.of_constr pars
else (* Let the usual code throw an error *) []
with Not_found -> []
else []
in
let app_f =
- match kind_of_term fj.uj_val with
+ match EConstr.kind !evdref fj.uj_val with
| Const (p, u) when Environ.is_projection p env.ExtraEnv.env ->
let p = Projection.make p false in
let pb = Environ.lookup_projection p env.ExtraEnv.env in
@@ -748,7 +752,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let argloc = loc_of_glob_constr c in
let resj = evd_comb1 (Coercion.inh_app_fun resolve_tc env.ExtraEnv.env) evdref resj in
let resty = whd_all env.ExtraEnv.env !evdref resj.uj_type in
- match kind_of_term resty with
+ match EConstr.kind !evdref resty with
| Prod (na,c1,c2) ->
let tycon = Some c1 in
let hj = pretype tycon env evdref lvar c in
@@ -768,18 +772,16 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let hj = pretype empty_tycon env evdref lvar c in
error_cant_apply_not_functional
~loc:(Loc.merge floc argloc) env.ExtraEnv.env !evdref
- resj [hj]
+ resj [|hj|]
in
let resj = apply_rec env 1 fj candargs args in
let resj =
- match evar_kind_of_term !evdref resj.uj_val with
+ match EConstr.kind !evdref resj.uj_val with
| App (f,args) ->
- let f = whd_evar !evdref f in
- if is_template_polymorphic env.ExtraEnv.env f then
+ if is_template_polymorphic env.ExtraEnv.env !evdref f then
(* Special case for inductive type applications that must be
refreshed right away. *)
- let sigma = !evdref in
- let c = mkApp (f,Array.map (whd_evar sigma) args) in
+ let c = mkApp (f, args) in
let c = evd_comb1 (Evarsolve.refresh_universes (Some true) env.ExtraEnv.env) evdref c in
let t = Retyping.get_type_of env.ExtraEnv.env !evdref c in
make_judge c (* use this for keeping evars: resj.uj_val *) t
@@ -805,7 +807,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
the substitution must also be applied on variables before they are
looked up in the rel context. *)
let var = LocalAssum (name, j.utj_val) in
- let j' = pretype rng (push_rel var env) evdref lvar c2 in
+ let j' = pretype rng (push_rel !evdref var env) evdref lvar c2 in
let name = ltac_interp_name lvar name in
let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in
inh_conv_coerce_to_tycon loc env evdref resj tycon
@@ -821,7 +823,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
{ j with utj_val = lift 1 j.utj_val }
| Name _ ->
let var = LocalAssum (name, j.utj_val) in
- let env' = push_rel var env in
+ let env' = push_rel !evdref var env in
pretype_type empty_valcon env' evdref lvar c2
in
let name = ltac_interp_name lvar name in
@@ -850,7 +852,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
looked up in the rel context. *)
let var = LocalDef (name, j.uj_val, t) in
let tycon = lift_tycon 1 tycon in
- let j' = pretype tycon (push_rel var env) evdref lvar c2 in
+ let j' = pretype tycon (push_rel !evdref var env) evdref lvar c2 in
let name = ltac_interp_name lvar name in
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
@@ -872,6 +874,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
user_err ~loc:loc (str "Destructing let on this type expects " ++
int cs.cs_nargs ++ str " variables.");
let fsign, record =
+ let set_name na d = set_name na (map_rel_decl EConstr.of_constr d) in
match get_projections env.ExtraEnv.env indf with
| None ->
List.map2 set_name (List.rev nal) cs.cs_args, false
@@ -879,6 +882,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let rec aux n k names l =
match names, l with
| na :: names, (LocalAssum (_,t) :: l) ->
+ let t = EConstr.of_constr t in
let proj = Projection.make ps.(cs.cs_nargs - k) true in
LocalDef (na, lift (cs.cs_nargs - n) (mkProj (proj, cj.uj_val)), t)
:: aux (n+1) (k + 1) names l
@@ -897,7 +901,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
mkCase (ci, p, cj.uj_val,[|f|])
else it_mkLambda_or_LetIn f fsign
in
- let env_f = push_rel_context fsign env in
+ let env_f = push_rel_context !evdref fsign env in
(* Make dependencies from arity signature impossible *)
let arsgn =
let arsgn,_ = get_arity env.ExtraEnv.env indf in
@@ -906,17 +910,18 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
else arsgn
in
let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in
+ let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let nar = List.length arsgn in
(match po with
| Some p ->
- let env_p = push_rel_context psign env in
+ let env_p = push_rel_context !evdref psign env in
let pj = pretype_type empty_valcon env_p evdref lvar p in
let ccl = nf_evar !evdref pj.utj_val in
- let psign = make_arity_signature env.ExtraEnv.env true indf in (* with names *)
+ let psign = make_arity_signature env.ExtraEnv.env !evdref true indf in (* with names *)
let p = it_mkLambda_or_LetIn ccl psign in
let inst =
- (Array.to_list cs.cs_concl_realargs)
- @[build_dependent_constructor cs] in
+ (Array.map_to_list EConstr.of_constr cs.cs_concl_realargs)
+ @[EConstr.of_constr (build_dependent_constructor cs)] in
let lp = lift cs.cs_nargs p in
let fty = hnf_lam_applist env.ExtraEnv.env !evdref lp inst in
let fj = pretype (mk_tycon fty) env_f evdref lvar d in
@@ -925,14 +930,14 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
Typing.check_allowed_sort env.ExtraEnv.env !evdref ind cj.uj_val p;
obj ind p cj.uj_val fj.uj_val
in
- { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
+ { uj_val = v; uj_type = (substl (realargs@[cj.uj_val]) ccl) }
| None ->
let tycon = lift_tycon cs.cs_nargs tycon in
let fj = pretype tycon env_f evdref lvar d in
let ccl = nf_evar !evdref fj.uj_type in
let ccl =
- if noccur_between 1 cs.cs_nargs ccl then
+ if noccur_between !evdref 1 cs.cs_nargs ccl then
lift (- cs.cs_nargs) ccl
else
error_cant_find_case_type ~loc env.ExtraEnv.env !evdref
@@ -966,19 +971,20 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
in
let nar = List.length arsgn in
let psign = LocalAssum (na, build_dependent_inductive env.ExtraEnv.env indf) :: arsgn in
+ let psign = List.map (fun d -> map_rel_decl EConstr.of_constr d) psign in
let pred,p = match po with
| Some p ->
- let env_p = push_rel_context psign env in
+ let env_p = push_rel_context !evdref psign env in
let pj = pretype_type empty_valcon env_p evdref lvar p in
let ccl = nf_evar !evdref pj.utj_val in
let pred = it_mkLambda_or_LetIn ccl psign in
- let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in
+ let typ = lift (- nar) (beta_applist !evdref (pred,[cj.uj_val])) in
pred, typ
| None ->
let p = match tycon with
| Some ty -> ty
| None ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
new_type_evar env evdref loc
in
it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in
@@ -987,18 +993,19 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let f cs b =
let n = Context.Rel.length cs.cs_args in
let pi = lift n pred in (* liftn n 2 pred ? *)
- let pi = beta_applist (pi, [build_dependent_constructor cs]) in
+ let pi = beta_applist !evdref (pi, [EConstr.of_constr (build_dependent_constructor cs)]) in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cs.cs_args in
let csgn =
if not !allow_anonymous_refs then
- List.map (set_name Anonymous) cs.cs_args
+ List.map (set_name Anonymous) cs_args
else
List.map (map_name (function Name _ as n -> n
| Anonymous -> Name Namegen.default_non_dependent_ident))
- cs.cs_args
+ cs_args
in
- let env_c = push_rel_context csgn env in
+ let env_c = push_rel_context !evdref csgn env in
let bj = pretype (mk_tycon pi) env_c evdref lvar b in
- it_mkLambda_or_LetIn bj.uj_val cs.cs_args in
+ it_mkLambda_or_LetIn bj.uj_val cs_args in
let b1 = f cstrs.(0) b1 in
let b2 = f cstrs.(1) b2 in
let v =
@@ -1013,7 +1020,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| GCases (loc,sty,po,tml,eqns) ->
Cases.compile_cases loc sty
- ((fun vtyc env evdref -> pretype vtyc (make_env env) evdref lvar),evdref)
+ ((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref)
tycon env.ExtraEnv.env (* loc *) (po,tml,eqns)
| GCast (loc,c,k) ->
@@ -1033,7 +1040,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
| VMcast ->
let cj = pretype empty_tycon env evdref lvar c in
let cty = nf_evar !evdref cj.uj_type and tval = nf_evar !evdref tval in
- if not (occur_existential cty || occur_existential tval) then
+ if not (occur_existential !evdref cty || occur_existential !evdref tval) then
let (evd,b) = Reductionops.vm_infer_conv env.ExtraEnv.env !evdref cty tval in
if b then (evdref := evd; cj, tval)
else
@@ -1061,7 +1068,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
let f decl (subst,update) =
let id = NamedDecl.get_id decl in
- let t = replace_vars subst (NamedDecl.get_type decl) in
+ let t = replace_vars subst (EConstr.of_constr (NamedDecl.get_type decl)) in
let c, update =
try
let c = List.assoc id update in
@@ -1087,13 +1094,22 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function
| GHole (loc, knd, naming, None) ->
+ let rec is_Type c = match EConstr.kind !evdref c with
+ | Sort s ->
+ begin match ESorts.kind !evdref s with
+ | Type _ -> true
+ | Prop _ -> false
+ end
+ | Cast (c, _, _) -> is_Type c
+ | _ -> false
+ in
(match valcon with
| Some v ->
let s =
let sigma = !evdref in
let t = Retyping.get_type_of env.ExtraEnv.env sigma v in
- match kind_of_term (whd_all env.ExtraEnv.env sigma t) with
- | Sort s -> s
+ match EConstr.kind sigma (whd_all env.ExtraEnv.env sigma t) with
+ | Sort s -> ESorts.kind sigma s
| Evar ev when is_Type (existential_type sigma ev) ->
evd_comb1 (define_evar_as_sort env.ExtraEnv.env) evdref ev
| _ -> anomaly (Pp.str "Found a type constraint which is not a type")
@@ -1101,7 +1117,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function
{ utj_val = v;
utj_type = s }
| None ->
- let env = ltac_interp_name_env k0 lvar env in
+ let env = ltac_interp_name_env k0 lvar env !evdref in
let s = evd_comb0 (new_sort_variable univ_flexible_alg) evdref in
{ utj_val = e_new_evar env evdref ~src:(loc, knd) ~naming (mkSort s);
utj_type = s})
@@ -1118,7 +1134,7 @@ and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function
~loc:(loc_of_glob_constr c) env.ExtraEnv.env !evdref tj.utj_val v
let ise_pretype_gen flags env sigma lvar kind c =
- let env = make_env env in
+ let env = make_env env sigma in
let evdref = ref sigma in
let k0 = Context.Rel.length (rel_context env) in
let c' = match kind with
@@ -1155,33 +1171,33 @@ let empty_lvar : ltac_var_map = {
ltac_genargs = Id.Map.empty;
}
-let on_judgment f j =
+let on_judgment sigma f j =
let c = mkCast(j.uj_val,DEFAULTcast, j.uj_type) in
- let (c,_,t) = destCast (f c) in
+ let (c,_,t) = destCast sigma (f c) in
{uj_val = c; uj_type = t}
let understand_judgment env sigma c =
- let env = make_env env in
+ let env = make_env env sigma in
let evdref = ref sigma in
let k0 = Context.Rel.length (rel_context env) in
let j = pretype k0 true empty_tycon env evdref empty_lvar c in
- let j = on_judgment (fun c ->
+ let j = on_judgment sigma (fun c ->
let evd, c = process_inference_flags all_and_fail_flags env.ExtraEnv.env sigma (!evdref,c) in
evdref := evd; c) j
in j, Evd.evar_universe_context !evdref
let understand_judgment_tcc env evdref c =
- let env = make_env env in
+ let env = make_env env !evdref in
let k0 = Context.Rel.length (rel_context env) in
let j = pretype k0 true empty_tycon env evdref empty_lvar c in
- on_judgment (fun c ->
+ on_judgment !evdref (fun c ->
let (evd,c) = process_inference_flags all_no_fail_flags env.ExtraEnv.env Evd.empty (!evdref,c) in
evdref := evd; c) j
let ise_pretype_gen_ctx flags env sigma lvar kind c =
let evd, c = ise_pretype_gen flags env sigma lvar kind c in
let evd, f = Evarutil.nf_evars_and_universes evd in
- f c, Evd.evar_universe_context evd
+ f (EConstr.Unsafe.to_constr c), Evd.evar_universe_context evd
(** Entry points of the high-level type synthesis algorithm *)
@@ -1192,7 +1208,8 @@ let understand
ise_pretype_gen_ctx flags env sigma empty_lvar expected_type c
let understand_tcc ?(flags=all_no_fail_flags) env sigma ?(expected_type=WithoutTypeConstraint) c =
- ise_pretype_gen flags env sigma empty_lvar expected_type c
+ let (sigma, c) = ise_pretype_gen flags env sigma empty_lvar expected_type c in
+ (sigma, c)
let understand_tcc_evars ?(flags=all_no_fail_flags) env evdref ?(expected_type=WithoutTypeConstraint) c =
let sigma, c = ise_pretype_gen flags env !evdref empty_lvar expected_type c in
@@ -1200,7 +1217,8 @@ let understand_tcc_evars ?(flags=all_no_fail_flags) env evdref ?(expected_type=W
c
let understand_ltac flags env sigma lvar kind c =
- ise_pretype_gen flags env sigma lvar kind c
+ let (sigma, c) = ise_pretype_gen flags env sigma lvar kind c in
+ (sigma, c)
let constr_flags = {
use_typeclasses = true;
@@ -1226,7 +1244,7 @@ let type_uconstr ?(flags = constr_flags)
end }
let pretype k0 resolve_tc typcon env evdref lvar t =
- pretype k0 resolve_tc typcon (make_env env) evdref lvar t
+ pretype k0 resolve_tc typcon (make_env env !evdref) evdref lvar t
let pretype_type k0 resolve_tc valcon env evdref lvar t =
- pretype_type k0 resolve_tc valcon (make_env env) evdref lvar t
+ pretype_type k0 resolve_tc valcon (make_env env !evdref) evdref lvar t
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 23957d557..f13c10b05 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -16,6 +16,7 @@ open Names
open Term
open Environ
open Evd
+open EConstr
open Glob_term
open Evarutil
open Misctypes
@@ -76,7 +77,7 @@ val allow_anonymous_refs : bool ref
evar_map is modified explicitly or by side-effect. *)
val understand_tcc : ?flags:inference_flags -> env -> evar_map ->
- ?expected_type:typing_constraint -> glob_constr -> open_constr
+ ?expected_type:typing_constraint -> glob_constr -> evar_map * constr
val understand_tcc_evars : ?flags:inference_flags -> env -> evar_map ref ->
?expected_type:typing_constraint -> glob_constr -> constr
@@ -105,7 +106,7 @@ val understand_ltac : inference_flags ->
unresolved evar remains, expanding evars. *)
val understand : ?flags:inference_flags -> ?expected_type:typing_constraint ->
- env -> evar_map -> glob_constr -> constr Evd.in_evar_universe_context
+ env -> evar_map -> glob_constr -> Constr.constr Evd.in_evar_universe_context
(** Idem but returns the judgment of the understood term *)
diff --git a/pretyping/program.ml b/pretyping/program.ml
index 4b6137b53..caa5a5c8a 100644
--- a/pretyping/program.ml
+++ b/pretyping/program.ml
@@ -29,6 +29,7 @@ let init_constant dir s () = coq_constant "Program" dir s
let init_reference dir s () = coq_reference "Program" dir s
let papp evdref r args =
+ let open EConstr in
let gr = delayed_force r in
mkApp (Evarutil.e_new_global evdref gr, args)
@@ -57,7 +58,9 @@ let coq_JMeq_refl = init_reference ["Logic";"JMeq"] "JMeq_refl"
let coq_not = init_constant ["Init";"Logic"] "not"
let coq_and = init_constant ["Init";"Logic"] "and"
-let mk_coq_not x = mkApp (delayed_force coq_not, [| x |])
+let delayed_force c = EConstr.of_constr (c ())
+
+let mk_coq_not x = EConstr.mkApp (delayed_force coq_not, [| x |])
let unsafe_fold_right f = function
hd :: tl -> List.fold_right f tl hd
@@ -67,7 +70,7 @@ let mk_coq_and l =
let and_typ = delayed_force coq_and in
unsafe_fold_right
(fun c conj ->
- mkApp (and_typ, [| c ; conj |]))
+ EConstr.mkApp (and_typ, [| c ; conj |]))
l
(* true = transparent by default, false = opaque if possible *)
diff --git a/pretyping/program.mli b/pretyping/program.mli
index 023ff8ca5..94a7bdcb6 100644
--- a/pretyping/program.mli
+++ b/pretyping/program.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open EConstr
open Globnames
(** A bunch of Coq constants used by Progam *)
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index cda052b79..bc9e3a1f4 100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -179,7 +179,7 @@ let cs_pattern_of_constr t =
with e when CErrors.noncritical e -> raise Not_found
end
| Rel n -> Default_cs, Some n, []
- | Prod (_,a,b) when not (Termops.dependent (mkRel 1) b) -> Prod_cs, None, [a; Termops.pop b]
+ | Prod (_,a,b) when Vars.noccurn 1 b -> Prod_cs, None, [a; Vars.lift (-1) b]
| Sort s -> Sort_cs (family_of_sort s), None, []
| _ ->
begin
@@ -191,7 +191,7 @@ let warn_projection_no_head_constant =
CWarnings.create ~name:"projection-no-head-constant" ~category:"typechecker"
(fun (t,con_pp,proji_sp_pp) ->
strbrk "Projection value has no head constant: "
- ++ Termops.print_constr t ++ strbrk " in canonical instance "
+ ++ Termops.print_constr (EConstr.of_constr t) ++ strbrk " in canonical instance "
++ con_pp ++ str " of " ++ proji_sp_pp ++ strbrk ", ignoring it.")
(* Intended to always succeed *)
@@ -202,8 +202,9 @@ let compute_canonical_projections warn (con,ind) =
let v = (mkConstU (con,u)) in
let ctx = Univ.ContextSet.of_context ctx in
let c = Environ.constant_value_in env (con,u) in
- let lt,t = Reductionops.splay_lam env Evd.empty c in
- let lt = List.rev_map snd lt in
+ let lt,t = Reductionops.splay_lam env Evd.empty (EConstr.of_constr c) in
+ let t = EConstr.Unsafe.to_constr t in
+ let lt = List.rev_map (snd %> EConstr.Unsafe.to_constr) lt in
let args = snd (decompose_app t) in
let { s_EXPECTEDPARAM = p; s_PROJ = lpj; s_PROJKIND = kl } =
lookup_structure ind in
@@ -255,8 +256,8 @@ let add_canonical_structure warn o =
in match ocs with
| None -> object_table := Refmap.add proj ((pat,s)::l) !object_table;
| Some (c, cs) ->
- let old_can_s = (Termops.print_constr cs.o_DEF)
- and new_can_s = (Termops.print_constr s.o_DEF) in
+ let old_can_s = (Termops.print_constr (EConstr.of_constr cs.o_DEF))
+ and new_can_s = (Termops.print_constr (EConstr.of_constr s.o_DEF)) in
let prj = (Nametab.pr_global_env Id.Set.empty proj)
and hd_val = (pr_cs_pattern cs_pat) in
if warn then warn_redundant_canonical_projection (hd_val,prj,new_can_s,old_can_s))
@@ -302,7 +303,8 @@ let check_and_decompose_canonical_structure ref =
let vc = match Environ.constant_opt_value_in env (sp, u) with
| Some vc -> vc
| None -> error_not_structure ref in
- let body = snd (splay_lam (Global.env()) Evd.empty vc) 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
| App (f,args) -> f,args
| _ -> error_not_structure ref in
@@ -322,14 +324,15 @@ let lookup_canonical_conversion (proj,pat) =
assoc_pat pat (Refmap.find proj !object_table)
let is_open_canonical_projection env sigma (c,args) =
+ let open EConstr in
try
- let ref = global_of_constr c in
+ let (ref, _) = Termops.global_of_constr sigma c in
let n = find_projection_nparams ref in
(** Check if there is some canonical projection attached to this structure *)
let _ = Refmap.find ref !object_table in
try
let arg = whd_all env sigma (Stack.nth args n) in
- let hd = match kind_of_term arg with App (hd, _) -> hd | _ -> arg in
- not (isConstruct hd)
+ let hd = match EConstr.kind sigma arg with App (hd, _) -> hd | _ -> arg in
+ not (isConstruct sigma hd)
with Failure _ -> false
with Not_found -> false
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index a6a90c751..7c0d0ec6d 100644
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -72,6 +72,6 @@ val pr_cs_pattern : cs_pattern -> Pp.std_ppcmds
val lookup_canonical_conversion : (global_reference * cs_pattern) -> constr * obj_typ
val declare_canonical_structure : global_reference -> unit
val is_open_canonical_projection :
- Environ.env -> Evd.evar_map -> (constr * constr Reductionops.Stack.t) -> bool
+ Environ.env -> Evd.evar_map -> Reductionops.state -> bool
val canonical_projections : unit ->
((global_reference * cs_pattern) * obj_typ) list
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index 1fdbbb412..270320538 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -10,11 +10,12 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Termops
open Univ
open Evd
open Environ
+open EConstr
+open Vars
open Context.Rel.Declaration
exception Elimconst
@@ -156,6 +157,7 @@ end
(** Machinery about stack of unfolded constants *)
module Cst_stack = struct
+ open EConstr
(** constant * params * args
- constant applied to params = term in head applied to args
@@ -191,24 +193,24 @@ module Cst_stack = struct
| (cst,params,[])::_ -> Some(cst,params)
| _ -> None
- let reference t = match best_cst t with
- | Some (c, _) when Term.isConst c -> Some (fst (Term.destConst c))
+ let reference sigma t = match best_cst t with
+ | Some (c, _) when isConst sigma c -> Some (fst (destConst sigma c))
| _ -> None
(** [best_replace d cst_l c] makes the best replacement for [d]
by [cst_l] in [c] *)
- let best_replace d cst_l c =
+ let best_replace sigma d cst_l c =
let reconstruct_head = List.fold_left
(fun t (i,args) -> mkApp (t,Array.sub args i (Array.length args - i))) in
List.fold_right
- (fun (cst,params,args) t -> Termops.replace_term
+ (fun (cst,params,args) t -> Termops.replace_term sigma
(reconstruct_head d args)
(applist (cst, List.rev params))
t) cst_l c
let pr l =
let open Pp in
- let p_c = Termops.print_constr in
+ let p_c c = Termops.print_constr c in
prlist_with_sep pr_semicolon
(fun (c,params,args) ->
hov 1 (str"(" ++ p_c c ++ str ")" ++ spc () ++ pr_sequence p_c params ++ spc () ++ str "(args:" ++
@@ -220,6 +222,7 @@ end
(** The type of (machine) stacks (= lambda-bar-calculus' contexts) *)
module Stack :
sig
+ open EConstr
type 'a app_node
val pr_app_node : ('a -> Pp.std_ppcmds) -> 'a app_node -> Pp.std_ppcmds
@@ -231,7 +234,7 @@ sig
| App of 'a app_node
| Case of case_info * 'a * 'a array * Cst_stack.t
| Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * 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
| Shift of int
| Update of 'a
@@ -242,10 +245,10 @@ sig
val append_app : 'a array -> 'a t -> 'a t
val decomp : 'a t -> ('a * 'a t) option
val decomp_node_last : 'a app_node -> 'a t -> ('a * 'a t)
- val equal : ('a * int -> 'a * int -> bool) -> (fixpoint * int -> fixpoint * int -> bool)
+ val equal : ('a * int -> 'a * int -> bool) -> (('a, 'a) pfixpoint * int -> ('a, 'a) pfixpoint * int -> bool)
-> 'a t -> 'a t -> (int * int) option
val compare_shape : 'a t -> 'a t -> bool
- val map : (constr -> constr) -> constr t -> constr t
+ val map : ('a -> 'a) -> 'a t -> 'a t
val fold2 : ('a -> constr -> constr -> 'a) -> 'a ->
constr t -> constr t -> 'a * int * int
val append_app_list : 'a list -> 'a t -> 'a t
@@ -258,10 +261,11 @@ sig
val args_size : 'a t -> int
val tail : int -> 'a t -> 'a t
val nth : 'a t -> int -> 'a
- val best_state : constr * constr t -> Cst_stack.t -> constr * constr t
- val zip : ?refold:bool -> constr * constr t -> constr
+ val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t
+ val zip : ?refold:bool -> evar_map -> constr * constr t -> constr
end =
struct
+ open EConstr
type 'a app_node = int * 'a array * int
(* first releavnt position, arguments, last relevant position *)
@@ -286,7 +290,7 @@ struct
| App of 'a app_node
| Case of Term.case_info * 'a * 'a array * Cst_stack.t
| Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * 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
| Shift of int
| Update of 'a
@@ -305,7 +309,7 @@ struct
str "ZProj(" ++ int n ++ pr_comma () ++ int m ++
pr_comma () ++ pr_con (Projection.constant p) ++ str ")"
| Fix (f,args,cst) ->
- str "ZFix(" ++ Termops.pr_fix Termops.print_constr f
+ str "ZFix(" ++ Termops.pr_fix pr_c f
++ pr_comma () ++ pr pr_c args ++ str ")"
| Cst (mem,curr,remains,params,cst_l) ->
str "ZCst(" ++ pr_cst_member pr_c mem ++ pr_comma () ++ int curr
@@ -533,11 +537,11 @@ struct
| None -> raise Not_found
(** This function breaks the abstraction of Cst_stack ! *)
- let best_state (_,sk as s) l =
+ let best_state sigma (_,sk as s) l =
let rec aux sk def = function
|(cst, params, []) -> (cst, append_app_list (List.rev params) sk)
|(cst, params, (i,t)::q) -> match decomp sk with
- | Some (el,sk') when Constr.equal el t.(i) ->
+ | Some (el,sk') when EConstr.eq_constr sigma el t.(i) ->
if i = pred (Array.length t)
then aux sk' def (cst, params, q)
else aux sk' def (cst, params, (succ i,t)::q)
@@ -546,59 +550,64 @@ struct
let constr_of_cst_member f sk =
match f with
- | Cst_const (c, u) -> mkConstU (c,u), sk
+ | Cst_const (c, u) -> mkConstU (c, EInstance.make u), sk
| Cst_proj p ->
match decomp sk with
| Some (hd, sk) -> mkProj (p, hd), sk
| None -> assert false
- let rec zip ?(refold=false) = function
+ let zip ?(refold=false) sigma s =
+ let rec zip = function
| f, [] -> f
| f, (App (i,a,j) :: s) ->
let a' = if Int.equal i 0 && Int.equal j (Array.length a - 1)
then a
else Array.sub a i (j - i + 1) in
- zip ~refold (mkApp (f, a'), s)
+ zip (mkApp (f, a'), s)
| f, (Case (ci,rt,br,cst_l)::s) when refold ->
- zip ~refold (best_state (mkCase (ci,rt,f,br), s) cst_l)
- | f, (Case (ci,rt,br,_)::s) -> zip ~refold (mkCase (ci,rt,f,br), s)
+ zip (best_state sigma (mkCase (ci,rt,f,br), s) cst_l)
+ | f, (Case (ci,rt,br,_)::s) -> zip (mkCase (ci,rt,f,br), s)
| f, (Fix (fix,st,cst_l)::s) when refold ->
- zip ~refold (best_state (mkFix fix, st @ (append_app [|f|] s)) cst_l)
- | f, (Fix (fix,st,_)::s) -> zip ~refold
+ zip (best_state sigma (mkFix fix, st @ (append_app [|f|] s)) cst_l)
+ | f, (Fix (fix,st,_)::s) -> zip
(mkFix fix, st @ (append_app [|f|] s))
| f, (Cst (cst,_,_,params,cst_l)::s) when refold ->
- zip ~refold (best_state (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l)
+ zip (best_state sigma (constr_of_cst_member cst (params @ (append_app [|f|] s))) cst_l)
| f, (Cst (cst,_,_,params,_)::s) ->
- zip ~refold (constr_of_cst_member cst (params @ (append_app [|f|] s)))
- | f, (Shift n::s) -> zip ~refold (lift n f, s)
+ zip (constr_of_cst_member cst (params @ (append_app [|f|] s)))
+ | f, (Shift n::s) -> zip (lift n f, s)
| f, (Proj (n,m,p,cst_l)::s) when refold ->
- zip ~refold (best_state (mkProj (p,f),s) cst_l)
- | f, (Proj (n,m,p,_)::s) -> zip ~refold (mkProj (p,f),s)
+ zip (best_state sigma (mkProj (p,f),s) cst_l)
+ | f, (Proj (n,m,p,_)::s) -> zip (mkProj (p,f),s)
| _, (Update _::_) -> assert false
+ in
+ zip s
+
end
(** The type of (machine) states (= lambda-bar-calculus' cuts) *)
type state = constr * constr Stack.t
-type contextual_reduction_function = env -> evar_map -> constr -> constr
-type reduction_function = contextual_reduction_function
+type contextual_reduction_function = env -> evar_map -> constr -> constr
+type reduction_function = contextual_reduction_function
type local_reduction_function = evar_map -> constr -> constr
type e_reduction_function = { e_redfun : 'r. env -> 'r Sigma.t -> constr -> (constr, 'r) Sigma.sigma }
-type contextual_stack_reduction_function =
- env -> evar_map -> constr -> constr * constr list
-type stack_reduction_function = contextual_stack_reduction_function
+type contextual_stack_reduction_function =
+ env -> evar_map -> constr -> constr * constr list
+type stack_reduction_function = contextual_stack_reduction_function
type local_stack_reduction_function =
evar_map -> constr -> constr * constr list
-type contextual_state_reduction_function =
- env -> evar_map -> state -> state
-type state_reduction_function = contextual_state_reduction_function
+type contextual_state_reduction_function =
+ env -> evar_map -> state -> state
+type state_reduction_function = contextual_state_reduction_function
type local_state_reduction_function = evar_map -> state -> state
let pr_state (tm,sk) =
let open Pp in
- h 0 (Termops.print_constr tm ++ str "|" ++ cut () ++ Stack.pr Termops.print_constr sk)
+ let pr c = Termops.print_constr c in
+ h 0 (pr tm ++ str "|" ++ cut () ++ Stack.pr pr sk)
(*************************************)
(*** Reduction Functions Operators ***)
@@ -612,16 +621,16 @@ let safe_meta_value sigma ev =
let strong whdfun env sigma t =
let rec strongrec env t =
- map_constr_with_full_binders push_rel strongrec env (whdfun env sigma t) in
+ map_constr_with_full_binders sigma push_rel strongrec env (whdfun env sigma t) in
strongrec env t
let local_strong whdfun sigma =
- let rec strongrec t = Constr.map strongrec (whdfun sigma t) in
+ let rec strongrec t = EConstr.map sigma strongrec (whdfun sigma t) in
strongrec
let rec strong_prodspine redfun sigma c =
let x = redfun sigma c in
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Prod (na,a,b) -> mkProd (na,a,strong_prodspine redfun sigma b)
| _ -> x
@@ -633,20 +642,25 @@ let eta = CClosure.RedFlags.mkflags [CClosure.RedFlags.fETA]
(* Beta Reduction tools *)
-let apply_subst recfun env refold cst_l t stack =
+let apply_subst recfun env sigma refold cst_l t stack =
let rec aux env cst_l t stack =
- match (Stack.decomp stack,kind_of_term t) with
+ match (Stack.decomp stack, EConstr.kind sigma t) with
| Some (h,stacktl), Lambda (_,_,c) ->
let cst_l' = if refold then Cst_stack.add_param h cst_l else cst_l in
aux (h::env) cst_l' c stacktl
- | _ -> recfun cst_l (substl env t, stack)
+ | _ -> recfun sigma cst_l (substl env t, stack)
in aux env cst_l t stack
-let stacklam recfun env t stack =
- apply_subst (fun _ -> recfun) env false Cst_stack.empty t stack
+let stacklam recfun env sigma t stack =
+ apply_subst (fun _ _ s -> recfun s) env sigma false Cst_stack.empty t stack
-let beta_applist (c,l) =
- stacklam Stack.zip [] c (Stack.append_app_list l Stack.empty)
+let beta_app sigma (c,l) =
+ let zip s = Stack.zip sigma s in
+ stacklam zip [] sigma c (Stack.append_app l Stack.empty)
+
+let beta_applist sigma (c,l) =
+ let zip s = Stack.zip sigma s in
+ stacklam zip [] sigma c (Stack.append_app_list l Stack.empty)
(* Iota reduction tools *)
@@ -657,7 +671,7 @@ type 'a miota_args = {
mcargs : 'a list; (* the constructor's arguments *)
mlf : 'a array } (* the branch code vector *)
-let reducible_mind_case c = match kind_of_term c with
+let reducible_mind_case sigma c = match EConstr.kind sigma c with
| Construct _ | CoFix _ -> true
| _ -> false
@@ -671,7 +685,7 @@ let reducible_mind_case c = match kind_of_term c with
f x := t. End M. Definition f := u. and say goodbye to any hope
of refolding M.f this way ...
*)
-let magicaly_constant_of_fixbody env reference bd = function
+let magicaly_constant_of_fixbody env sigma reference bd = function
| Name.Anonymous -> bd
| Name.Name id ->
try
@@ -681,7 +695,7 @@ let magicaly_constant_of_fixbody env reference bd = function
match constant_opt_value_in env (cst,u) with
| None -> bd
| Some t ->
- let csts = Universes.eq_constr_universes t bd in
+ let csts = EConstr.eq_constr_universes sigma (EConstr.of_constr t) bd in
begin match csts with
| Some csts ->
let subst = Universes.Constraints.fold (fun (l,d,r) acc ->
@@ -689,13 +703,13 @@ let magicaly_constant_of_fixbody env reference bd = function
csts Univ.LMap.empty
in
let inst = Instance.subst_fn (fun u -> Univ.LMap.find u subst) u in
- mkConstU (cst,inst)
+ mkConstU (cst, EInstance.make inst)
| None -> bd
end
with
| Not_found -> bd
-let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies)) =
+let contract_cofix ?env sigma ?reference (bodynum,(names,types,bodies as typedbodies)) =
let nbodies = Array.length bodies in
let make_Fi j =
let ind = nbodies-j-1 in
@@ -707,37 +721,37 @@ let contract_cofix ?env ?reference (bodynum,(names,types,bodies as typedbodies))
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
(** Similar to the "fix" case below *)
-let reduce_and_refold_cofix recfun env refold cst_l cofix sk =
+let reduce_and_refold_cofix recfun env sigma refold cst_l cofix sk =
let raw_answer =
let env = if refold then Some env else None in
- contract_cofix ?env ?reference:(Cst_stack.reference cst_l) cofix in
+ contract_cofix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) cofix in
apply_subst
- (fun x (t,sk') ->
+ (fun sigma x (t,sk') ->
let t' =
- if refold then Cst_stack.best_replace (mkCoFix cofix) cst_l t else t in
+ if refold then Cst_stack.best_replace sigma (mkCoFix cofix) cst_l t else t in
recfun x (t',sk'))
- [] refold Cst_stack.empty raw_answer sk
+ [] sigma refold Cst_stack.empty raw_answer sk
-let reduce_mind_case mia =
- match kind_of_term mia.mconstr with
+let reduce_mind_case sigma mia =
+ match EConstr.kind sigma mia.mconstr with
| Construct ((ind_sp,i),u) ->
(* let ncargs = (fst mia.mci).(i-1) in*)
let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1),real_cargs)
| CoFix cofix ->
- let cofix_def = contract_cofix cofix in
+ let cofix_def = contract_cofix sigma cofix in
mkCase (mia.mci, mia.mP, applist(cofix_def,mia.mcargs), mia.mlf)
| _ -> assert false
(* contracts fix==FIX[nl;i](A1...Ak;[F1...Fk]{B1....Bk}) to produce
Bi[Fj --> FIX[nl;j](A1...Ak;[F1...Fk]{B1...Bk})] *)
-let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) =
+let contract_fix ?env sigma ?reference ((recindices,bodynum),(names,types,bodies as typedbodies)) =
let nbodies = Array.length recindices in
let make_Fi j =
let ind = nbodies-j-1 in
@@ -749,7 +763,7 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty
| Some e ->
match reference with
| None -> bd
- | Some r -> magicaly_constant_of_fixbody e r bd names.(ind) in
+ | Some r -> magicaly_constant_of_fixbody e sigma r bd names.(ind) in
let closure = List.init nbodies make_Fi in
substl closure bodies.(bodynum)
@@ -757,18 +771,18 @@ let contract_fix ?env ?reference ((recindices,bodynum),(names,types,bodies as ty
replace the fixpoint by the best constant from [cst_l]
Other rels are directly substituted by constants "magically found from the
context" in contract_fix *)
-let reduce_and_refold_fix recfun env refold cst_l fix sk =
+let reduce_and_refold_fix recfun env sigma refold cst_l fix sk =
let raw_answer =
let env = if refold then None else Some env in
- contract_fix ?env ?reference:(Cst_stack.reference cst_l) fix in
+ contract_fix ?env sigma ?reference:(Cst_stack.reference sigma cst_l) fix in
apply_subst
- (fun x (t,sk') ->
+ (fun sigma x (t,sk') ->
let t' =
if refold then
- Cst_stack.best_replace (mkFix fix) cst_l t
+ Cst_stack.best_replace sigma (mkFix fix) cst_l t
else t
in recfun x (t',sk'))
- [] refold Cst_stack.empty raw_answer sk
+ [] sigma refold Cst_stack.empty raw_answer sk
let fix_recarg ((recindices,bodynum),_) stack =
assert (0 <= bodynum && bodynum < Array.length recindices);
@@ -801,30 +815,32 @@ let _ = Goptions.declare_bool_option {
Goptions.optwrite = (fun a -> debug_RAKAM:=a);
}
-let equal_stacks (x, l) (y, l') =
- let f_equal (x,lft1) (y,lft2) = Constr.equal (Vars.lift lft1 x) (Vars.lift lft2 y) in
- let eq_fix (a,b) (c,d) = f_equal (Constr.mkFix a, b) (Constr.mkFix c, d) in
+let equal_stacks sigma (x, l) (y, l') =
+ let f_equal (x,lft1) (y,lft2) = eq_constr sigma (Vars.lift lft1 x) (Vars.lift lft2 y) in
+ let eq_fix (a,b) (c,d) = f_equal (mkFix a, b) (mkFix c, d) in
match Stack.equal f_equal eq_fix l l' with
| None -> false
| Some (lft1,lft2) -> f_equal (x, lft1) (y, lft2)
let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
let open Context.Named.Declaration in
- let rec whrec cst_l (x, stack as s) =
+ let rec whrec cst_l (x, stack) =
let () = if !debug_RAKAM then
let open Pp in
+ let pr c = Termops.print_constr c in
Feedback.msg_notice
- (h 0 (str "<<" ++ Termops.print_constr x ++
+ (h 0 (str "<<" ++ pr x ++
str "|" ++ cut () ++ Cst_stack.pr cst_l ++
- str "|" ++ cut () ++ Stack.pr Termops.print_constr stack ++
+ str "|" ++ cut () ++ Stack.pr pr stack ++
str ">>"))
in
+ let c0 = EConstr.kind sigma x in
let fold () =
let () = if !debug_RAKAM then
let open Pp in Feedback.msg_notice (str "<><><><><>") in
- (s,cst_l)
+ ((EConstr.of_kind c0, stack),cst_l)
in
- match kind_of_term x with
+ match c0 with
| Rel n when CClosure.RedFlags.red_set flags CClosure.RedFlags.fDELTA ->
(match lookup_rel n env with
| LocalDef (_,body,_) -> whrec Cst_stack.empty (lift n body, stack)
@@ -834,18 +850,17 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| LocalDef (_,body,_) ->
whrec (if refold then Cst_stack.add_cst (mkVar id) cst_l else cst_l) (body, stack)
| _ -> fold ())
- | Evar ev ->
- (match safe_evar_value sigma ev with
- | Some body -> whrec cst_l (body, stack)
- | None -> fold ())
+ | Evar ev -> fold ()
| Meta ev ->
(match safe_meta_value sigma ev with
- | Some body -> whrec cst_l (body, stack)
+ | Some body -> whrec cst_l (EConstr.of_constr body, stack)
| None -> fold ())
| Const (c,u as const) when CClosure.RedFlags.red_set flags (CClosure.RedFlags.fCONST c) ->
- (match constant_opt_value_in env const with
+ let u' = EInstance.kind sigma u in
+ (match constant_opt_value_in env (fst const, u') with
| None -> fold ()
| Some body ->
+ let body = EConstr.of_constr body in
if not tactic_mode
then whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l)
(body, stack)
@@ -862,12 +877,12 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
let (tm',sk'),cst_l' =
whrec (Cst_stack.add_cst (mkConstU const) cst_l) (body, app_sk)
in
- let rec is_case x = match kind_of_term x with
+ let rec is_case x = match EConstr.kind sigma x with
| Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
| App (hd, _) -> is_case hd
| Case _ -> true
| _ -> false in
- if equal_stacks (x, app_sk) (tm', sk')
+ if equal_stacks sigma (x, app_sk) (tm', sk')
|| Stack.will_expose_iota sk'
|| is_case tm'
then fold ()
@@ -881,7 +896,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| None -> fold ()
| Some (bef,arg,s') ->
whrec Cst_stack.empty
- (arg,Stack.Cst(Stack.Cst_const const,curr,remains,bef,cst_l)::s')
+ (arg,Stack.Cst(Stack.Cst_const (fst const, u'),curr,remains,bef,cst_l)::s')
)
| Proj (p, c) when CClosure.RedFlags.red_projection flags p ->
(let pb = lookup_projection p env in
@@ -895,7 +910,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| None ->
let stack' = (c, Stack.Proj (npars, arg, p, cst_l) :: stack) in
let stack'', csts = whrec Cst_stack.empty stack' in
- if equal_stacks stack' stack'' then fold ()
+ if equal_stacks sigma stack' stack'' then fold ()
else stack'', csts
| Some (recargs, nargs, flags) ->
if (List.mem `ReductionNeverUnfold flags
@@ -925,7 +940,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
Stack.append_app [|c|] bef,cst_l)::s'))
| LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA ->
- apply_subst whrec [b] refold cst_l c stack
+ apply_subst (fun _ -> whrec) [b] sigma refold cst_l c stack
| Cast (c,_,_) -> whrec cst_l (c, stack)
| App (f,cl) ->
whrec
@@ -934,20 +949,20 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
| Lambda (na,t,c) ->
(match Stack.decomp stack with
| Some _ when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA ->
- apply_subst whrec [] refold cst_l x stack
+ apply_subst (fun _ -> whrec) [] sigma refold cst_l x stack
| None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA ->
- let env' = push_rel (LocalAssum (na,t)) env in
+ let env' = push_rel (LocalAssum (na, t)) env in
let whrec' = whd_state_gen ~refold ~tactic_mode flags env' sigma in
- (match kind_of_term (Stack.zip ~refold (fst (whrec' (c, Stack.empty)))) with
+ (match EConstr.kind sigma (Stack.zip ~refold sigma (fst (whrec' (c, Stack.empty)))) with
| App (f,cl) ->
let napp = Array.length cl in
if napp > 0 then
let (x', l'),_ = whrec' (Array.last cl, Stack.empty) in
- match kind_of_term x', l' with
+ match EConstr.kind sigma x', l' with
| Rel 1, [] ->
let lc = Array.sub cl 0 (napp-1) in
- let u = if Int.equal napp 1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,Stack.empty),Cst_stack.empty else fold ()
+ let u = if Int.equal napp 1 then f else mkApp (f,lc) in
+ if noccurn sigma 1 u then (pop u,Stack.empty),Cst_stack.empty else fold ()
| _ -> fold ()
else fold ()
| _ -> fold ())
@@ -972,11 +987,11 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
|args, (Stack.Proj (n,m,p,_)::s') when use_match ->
whrec Cst_stack.empty (Stack.nth args (n+m), s')
|args, (Stack.Fix (f,s',cst_l)::s'') when use_fix ->
- let x' = Stack.zip(x,args) in
+ let x' = Stack.zip sigma (x, args) in
let out_sk = s' @ (Stack.append_app [|x'|] s'') in
- reduce_and_refold_fix whrec env refold cst_l f out_sk
+ reduce_and_refold_fix whrec env sigma refold cst_l f out_sk
|args, (Stack.Cst (const,curr,remains,s',cst_l) :: s'') ->
- let x' = Stack.zip(x,args) in
+ let x' = Stack.zip sigma (x, args) in
begin match remains with
| [] ->
(match const with
@@ -984,6 +999,8 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
(match constant_opt_value_in env const with
| None -> fold ()
| Some body ->
+ let const = (fst const, EInstance.make (snd const)) in
+ let body = EConstr.of_constr body in
whrec (if refold then Cst_stack.add_cst (mkConstU const) cst_l else cst_l)
(body, s' @ (Stack.append_app [|x'|] s'')))
| Stack.Cst_proj p ->
@@ -1010,7 +1027,7 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then
match Stack.strip_app stack with
|args, ((Stack.Case _ |Stack.Proj _)::s') ->
- reduce_and_refold_cofix whrec env refold cst_l cofix stack
+ reduce_and_refold_cofix whrec env sigma refold cst_l cofix stack
|_ -> fold ()
else fold ()
@@ -1019,31 +1036,33 @@ let rec whd_state_gen ?csts ~refold ~tactic_mode flags env sigma =
in
fun xs ->
let (s,cst_l as res) = whrec (Option.default Cst_stack.empty csts) xs in
- if tactic_mode then (Stack.best_state s cst_l,Cst_stack.empty) else res
+ if tactic_mode then (Stack.best_state sigma s cst_l,Cst_stack.empty) else res
(** reduction machine without global env and refold machinery *)
let local_whd_state_gen flags sigma =
- let rec whrec (x, stack as s) =
- match kind_of_term x with
+ let rec whrec (x, stack) =
+ let c0 = EConstr.kind sigma x in
+ let s = (EConstr.of_kind c0, stack) in
+ match c0 with
| LetIn (_,b,_,c) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fZETA ->
- stacklam whrec [b] c stack
+ stacklam whrec [b] sigma c stack
| Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, Stack.append_app cl stack)
| Lambda (_,_,c) ->
(match Stack.decomp stack with
| Some (a,m) when CClosure.RedFlags.red_set flags CClosure.RedFlags.fBETA ->
- stacklam whrec [a] c m
+ stacklam whrec [a] sigma c m
| None when CClosure.RedFlags.red_set flags CClosure.RedFlags.fETA ->
- (match kind_of_term (Stack.zip (whrec (c, Stack.empty))) with
+ (match EConstr.kind sigma (Stack.zip sigma (whrec (c, Stack.empty))) with
| App (f,cl) ->
let napp = Array.length cl in
if napp > 0 then
let x', l' = whrec (Array.last cl, Stack.empty) in
- match kind_of_term x', l' with
+ match EConstr.kind sigma x', l' with
| Rel 1, [] ->
let lc = Array.sub cl 0 (napp-1) in
- let u = if Int.equal napp 1 then f else appvect (f,lc) in
- if noccurn 1 u then (pop u,Stack.empty) else s
+ let u = if Int.equal napp 1 then f else mkApp (f,lc) in
+ if noccurn sigma 1 u then (pop u,Stack.empty) else s
| _ -> s
else s
| _ -> s)
@@ -1063,14 +1082,10 @@ let local_whd_state_gen flags sigma =
|None -> s
|Some (bef,arg,s') -> whrec (arg, Stack.Fix(f,bef,Cst_stack.empty)::s'))
- | Evar ev ->
- (match safe_evar_value sigma ev with
- Some c -> whrec (c,stack)
- | None -> s)
-
+ | Evar ev -> s
| Meta ev ->
(match safe_meta_value sigma ev with
- Some c -> whrec (c,stack)
+ Some c -> whrec (EConstr.of_constr c,stack)
| None -> s)
| Construct ((ind,c),u) ->
@@ -1083,8 +1098,8 @@ let local_whd_state_gen flags sigma =
|args, (Stack.Proj (n,m,p,_) :: s') when use_match ->
whrec (Stack.nth args (n+m), s')
|args, (Stack.Fix (f,s',cst)::s'') when use_fix ->
- let x' = Stack.zip(x,args) in
- whrec (contract_fix f, s' @ (Stack.append_app [|x'|] s''))
+ let x' = Stack.zip sigma (x,args) in
+ whrec (contract_fix sigma f, s' @ (Stack.append_app [|x'|] s''))
|_, (Stack.App _|Stack.Update _|Stack.Shift _|Stack.Cst _)::_ -> assert false
|_, _ -> s
else s
@@ -1093,7 +1108,7 @@ let local_whd_state_gen flags sigma =
if CClosure.RedFlags.red_set flags CClosure.RedFlags.fCOFIX then
match Stack.strip_app stack with
|args, ((Stack.Case _ | Stack.Proj _)::s') ->
- whrec (contract_cofix cofix, stack)
+ whrec (contract_cofix sigma cofix, stack)
|_ -> s
else s
@@ -1106,7 +1121,7 @@ let raw_whd_state_gen flags env =
f
let stack_red_of_state_red f =
- let f sigma x = decompose_app (Stack.zip (f sigma (x, Stack.empty))) in
+ let f sigma x = EConstr.decompose_app sigma (Stack.zip sigma (f sigma (x, Stack.empty))) in
f
(* Drops the Cst_stack *)
@@ -1114,11 +1129,11 @@ let iterate_whd_gen refold flags env sigma s =
let rec aux t =
let (hd,sk),_ = whd_state_gen refold false flags env sigma (t,Stack.empty) in
let whd_sk = Stack.map aux sk in
- Stack.zip ~refold (hd,whd_sk)
+ Stack.zip sigma ~refold (hd,whd_sk)
in aux s
let red_of_state_red f sigma x =
- Stack.zip (f sigma (x,Stack.empty))
+ Stack.zip sigma (f sigma (x,Stack.empty))
(* 0. No Reduction Functions *)
@@ -1173,7 +1188,7 @@ let whd_allnolet env =
(* 4. Ad-hoc eta reduction, does not subsitute evars *)
-let shrink_eta c = Stack.zip (local_whd_state_gen eta Evd.empty (c,Stack.empty))
+let shrink_eta c = Stack.zip Evd.empty (local_whd_state_gen eta Evd.empty (c,Stack.empty))
(* 5. Zeta Reduction Functions *)
@@ -1195,9 +1210,17 @@ let nf_evar = Evarutil.nf_evar
let clos_norm_flags flgs env sigma t =
try
let evars ev = safe_evar_value sigma ev in
- CClosure.norm_val
+ EConstr.of_constr (CClosure.norm_val
(CClosure.create_clos_infos ~evars flgs env)
- (CClosure.inject t)
+ (CClosure.inject (EConstr.Unsafe.to_constr t)))
+ with e when is_anomaly e -> error "Tried to normalize ill-typed term"
+
+let clos_whd_flags flgs env sigma t =
+ try
+ let evars ev = safe_evar_value sigma ev in
+ EConstr.of_constr (CClosure.whd_val
+ (CClosure.create_clos_infos ~evars flgs env)
+ (CClosure.inject (EConstr.Unsafe.to_constr t)))
with e when is_anomaly e -> error "Tried to normalize ill-typed term"
let nf_beta = clos_norm_flags CClosure.beta (Global.env ())
@@ -1238,6 +1261,14 @@ let report_anomaly _ =
let e = CErrors.push e in
iraise e
+let f_conv ?l2r ?reds env ?evars x y =
+ let inj = EConstr.Unsafe.to_constr in
+ Reduction.conv ?l2r ?reds env ?evars (inj x) (inj y)
+
+let f_conv_leq ?l2r ?reds env ?evars x y =
+ let inj = EConstr.Unsafe.to_constr in
+ Reduction.conv_leq ?l2r ?reds env ?evars (inj x) (inj y)
+
let test_trans_conversion (f: constr Reduction.extended_conversion_function) reds env sigma x y =
try
let evars ev = safe_evar_value sigma ev in
@@ -1246,16 +1277,16 @@ let test_trans_conversion (f: constr Reduction.extended_conversion_function) red
with Reduction.NotConvertible -> false
| e when is_anomaly e -> report_anomaly e
-let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv reds env sigma
-let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion Reduction.conv_leq reds env sigma
+let is_conv ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv reds env sigma
+let is_conv_leq ?(reds=full_transparent_state) env sigma = test_trans_conversion f_conv_leq reds env sigma
let is_fconv ?(reds=full_transparent_state) = function
| Reduction.CONV -> is_conv ~reds
| Reduction.CUMUL -> is_conv_leq ~reds
let check_conv ?(pb=Reduction.CUMUL) ?(ts=full_transparent_state) env sigma x y =
let f = match pb with
- | Reduction.CONV -> Reduction.conv
- | Reduction.CUMUL -> Reduction.conv_leq
+ | Reduction.CONV -> f_conv
+ | Reduction.CUMUL -> f_conv_leq
in
try f ~reds:ts env ~evars:(safe_evar_value sigma, Evd.universes sigma) x y; true
with Reduction.NotConvertible -> false
@@ -1279,6 +1310,9 @@ let sigma_univ_state =
let infer_conv_gen conv_fun ?(catch_incon=true) ?(pb=Reduction.CUMUL)
?(ts=full_transparent_state) env sigma x y =
+ (** FIXME *)
+ let x = EConstr.Unsafe.to_constr x in
+ let y = EConstr.Unsafe.to_constr y in
try
let fold cstr sigma =
try Some (Evd.add_universe_constraints sigma cstr)
@@ -1319,37 +1353,37 @@ let vm_infer_conv ?(pb=Reduction.CUMUL) env t1 t2 =
(* Special-Purpose Reduction *)
(********************************************************************)
-let whd_meta sigma c = match kind_of_term c with
- | Meta p -> (try meta_value sigma p with Not_found -> c)
+let whd_meta sigma c = match EConstr.kind sigma c with
+ | Meta p -> (try EConstr.of_constr (meta_value sigma p) with Not_found -> c)
| _ -> c
let default_plain_instance_ident = Id.of_string "H"
(* Try to replace all metas. Does not replace metas in the metas' values
* Differs from (strong whd_meta). *)
-let plain_instance s c =
- let rec irec n u = match kind_of_term u with
+let plain_instance sigma s c =
+ let rec irec n u = match EConstr.kind sigma u with
| Meta p -> (try lift n (Metamap.find p s) with Not_found -> u)
- | App (f,l) when isCast f ->
- let (f,_,t) = destCast f in
+ | App (f,l) when isCast sigma f ->
+ let (f,_,t) = destCast sigma f in
let l' = CArray.Fun1.smartmap irec n l in
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Meta p ->
(* Don't flatten application nodes: this is used to extract a
proof-term from a proof-tree and we want to keep the structure
of the proof-tree *)
(try let g = Metamap.find p s in
- match kind_of_term g with
+ match EConstr.kind sigma g with
| App _ ->
let l' = CArray.Fun1.smartmap lift 1 l' in
mkLetIn (Name default_plain_instance_ident,g,t,mkApp(mkRel 1, l'))
| _ -> mkApp (g,l')
with Not_found -> mkApp (f,l'))
| _ -> mkApp (irec n f,l'))
- | Cast (m,_,_) when isMeta m ->
- (try lift n (Metamap.find (destMeta m) s) with Not_found -> u)
+ | Cast (m,_,_) when isMeta sigma m ->
+ (try lift n (Metamap.find (destMeta sigma m) s) with Not_found -> u)
| _ ->
- map_constr_with_binders succ irec n u
+ map_with_binders sigma succ irec n u
in
if Metamap.is_empty s then c
else irec 0 c
@@ -1390,7 +1424,7 @@ let plain_instance s c =
let instance sigma s c =
(* if s = [] then c else *)
- local_strong whd_betaiota sigma (plain_instance s c)
+ local_strong whd_betaiota sigma (plain_instance sigma s c)
(* pseudo-reduction rule:
* [hnf_prod_app env s (Prod(_,B)) N --> B[N]
@@ -1399,34 +1433,37 @@ let instance sigma s c =
* error message. *)
let hnf_prod_app env sigma t n =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (_,_,b) -> subst1 n b
| _ -> anomaly ~label:"hnf_prod_app" (Pp.str "Need a product")
let hnf_prod_appvect env sigma t nl =
- Array.fold_left (hnf_prod_app env sigma) t nl
+ Array.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl
let hnf_prod_applist env sigma t nl =
- List.fold_left (hnf_prod_app env sigma) t nl
+ List.fold_left (fun acc t -> hnf_prod_app env sigma acc t) t nl
let hnf_lam_app env sigma t n =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Lambda (_,_,b) -> subst1 n b
| _ -> anomaly ~label:"hnf_lam_app" (Pp.str "Need an abstraction")
let hnf_lam_appvect env sigma t nl =
- Array.fold_left (hnf_lam_app env sigma) t nl
+ Array.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl
let hnf_lam_applist env sigma t nl =
- List.fold_left (hnf_lam_app env sigma) t nl
+ List.fold_left (fun acc t -> hnf_lam_app env sigma acc t) t nl
+
+let bind_assum (na, t) =
+ (na, t)
let splay_prod env sigma =
let rec decrec env m c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
- ((n,a)::m) c0
+ (bind_assum (n,a)::m) c0
| _ -> m,t
in
decrec env []
@@ -1434,10 +1471,10 @@ let splay_prod env sigma =
let splay_lam env sigma =
let rec decrec env m c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Lambda (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
- ((n,a)::m) c0
+ (bind_assum (n,a)::m) c0
| _ -> m,t
in
decrec env []
@@ -1445,7 +1482,7 @@ let splay_lam env sigma =
let splay_prod_assum env sigma =
let rec prodec_rec env l c =
let t = whd_allnolet env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (x,t,c) ->
prodec_rec (push_rel (LocalAssum (x,t)) env)
(Context.Rel.add (LocalAssum (x,t)) l) c
@@ -1455,14 +1492,14 @@ let splay_prod_assum env sigma =
| Cast (c,_,_) -> prodec_rec env l c
| _ ->
let t' = whd_all env sigma t in
- if Term.eq_constr t t' then l,t
+ if EConstr.eq_constr sigma t t' then l,t
else prodec_rec env l t'
in
prodec_rec env Context.Rel.empty
let splay_arity env sigma c =
let l, c = splay_prod env sigma c in
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Sort s -> l,s
| _ -> invalid_arg "splay_arity"
@@ -1470,7 +1507,7 @@ let sort_of_arity env sigma c = snd (splay_arity env sigma c)
let splay_prod_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
- match kind_of_term (whd_all env sigma c) with
+ match EConstr.kind sigma (whd_all env sigma c) with
| Prod (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
(m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
@@ -1480,7 +1517,7 @@ let splay_prod_n env sigma n =
let splay_lam_n env sigma n =
let rec decrec env m ln c = if Int.equal m 0 then (ln,c) else
- match kind_of_term (whd_all env sigma c) with
+ match EConstr.kind sigma (whd_all env sigma c) with
| Lambda (n,a,c0) ->
decrec (push_rel (LocalAssum (n,a)) env)
(m-1) (Context.Rel.add (LocalAssum (n,a)) ln) c0
@@ -1489,7 +1526,7 @@ let splay_lam_n env sigma n =
decrec env n Context.Rel.empty
let is_sort env sigma t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Sort s -> true
| _ -> false
@@ -1505,15 +1542,15 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s =
|args, (Stack.Case _ :: _ as stack') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if reducible_mind_case t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
+ if reducible_mind_case sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
|args, (Stack.Fix _ :: _ as stack') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if isConstruct t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
+ if isConstruct sigma t_o then whrec csts_o (t_o, stack_o@stack') else s,csts'
|args, (Stack.Proj (n,m,p,_) :: stack'') ->
let (t_o,stack_o),csts_o = whd_state_gen ~csts:csts' ~refold ~tactic_mode
(CClosure.RedFlags.red_add_transparent CClosure.all ts) env sigma (t,args) in
- if isConstruct t_o then
+ if isConstruct sigma t_o then
whrec Cst_stack.empty (Stack.nth stack_o (n+m), stack'')
else s,csts'
|_, ((Stack.App _| Stack.Shift _|Stack.Update _|Stack.Cst _) :: _|[]) -> s,csts'
@@ -1522,7 +1559,7 @@ let whd_betaiota_deltazeta_for_iota_state ts env sigma csts s =
let find_conclusion env sigma =
let rec decrec env c =
let t = whd_all env sigma c in
- match kind_of_term t with
+ match EConstr.kind sigma t with
| Prod (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0
| Lambda (x,t,c0) -> decrec (push_rel (LocalAssum (x,t)) env) c0
| t -> t
@@ -1542,7 +1579,7 @@ let meta_value evd mv =
match meta_opt_fvalue evd mv with
| Some (b,_) ->
let metas = Metamap.bind valrec b.freemetas in
- instance evd metas b.rebus
+ instance evd metas (EConstr.of_constr b.rebus)
| None -> mkMeta mv
in
valrec mv
@@ -1554,7 +1591,10 @@ let meta_instance sigma b =
let c_sigma = Metamap.bind (fun mv -> meta_value sigma mv) fm in
instance sigma c_sigma b.rebus
-let nf_meta sigma c = meta_instance sigma (mk_freelisted c)
+let nf_meta sigma c =
+ let c = EConstr.Unsafe.to_constr c in
+ let cl = mk_freelisted c in
+ meta_instance sigma { cl with rebus = EConstr.of_constr cl.rebus }
(* Instantiate metas that create beta/iota redexes *)
@@ -1568,78 +1608,80 @@ let meta_reducible_instance evd b =
in
let metas = Metaset.fold fold fm Metamap.empty in
let rec irec u =
- let u = whd_betaiota Evd.empty u in
- match kind_of_term u with
- | Case (ci,p,c,bl) when isMeta (strip_outer_cast c) ->
- let m = destMeta (strip_outer_cast c) in
+ let u = whd_betaiota Evd.empty u (** FIXME *) in
+ match EConstr.kind evd u with
+ | Case (ci,p,c,bl) when EConstr.isMeta evd (strip_outer_cast evd c) ->
+ let m = destMeta evd (strip_outer_cast evd c) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isConstruct g || not is_coerce then Some g else None
+ if isConstruct evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkCase (ci,p,g,bl))
| None -> mkCase (ci,irec p,c,Array.map irec bl))
- | App (f,l) when isMeta (strip_outer_cast f) ->
- let m = destMeta (strip_outer_cast f) in
+ | App (f,l) when EConstr.isMeta evd (strip_outer_cast evd f) ->
+ let m = destMeta evd (strip_outer_cast evd f) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isLambda g || not is_coerce then Some g else None
+ if isLambda evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkApp (g,l))
| None -> mkApp (f,Array.map irec l))
| Meta m ->
(try let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
if not is_coerce then irec g else u
with Not_found -> u)
- | Proj (p,c) when isMeta c || isCast c && isMeta (pi1 (destCast c)) ->
- let m = try destMeta c with _ -> destMeta (pi1 (destCast c)) in
+ | Proj (p,c) when isMeta evd c || isCast evd c && isMeta evd (pi1 (destCast evd c)) ->
+ let m = try destMeta evd c with _ -> destMeta evd (pi1 (destCast evd c)) in
(match
try
let g, s = Metamap.find m metas in
+ let g = EConstr.of_constr g in
let is_coerce = match s with CoerceToType -> true | _ -> false in
- if isConstruct g || not is_coerce then Some g else None
+ if isConstruct evd g || not is_coerce then Some g else None
with Not_found -> None
with
| Some g -> irec (mkProj (p,g))
| None -> mkProj (p,c))
- | _ -> Constr.map irec u
+ | _ -> EConstr.map evd irec u
in
if Metaset.is_empty fm then (* nf_betaiota? *) b.rebus
else irec b.rebus
-let head_unfold_under_prod ts env _ c =
- let unfold (cst,u as cstu) =
+let head_unfold_under_prod ts env sigma c =
+ let unfold (cst,u) =
+ let cstu = (cst, EInstance.kind sigma u) in
if Cpred.mem cst (snd ts) then
match constant_opt_value_in env cstu with
- | Some c -> c
- | None -> mkConstU cstu
- else mkConstU cstu in
+ | Some c -> EConstr.of_constr c
+ | None -> mkConstU (cst, u)
+ else mkConstU (cst, u) in
let rec aux c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Prod (n,t,c) -> mkProd (n,aux t, aux c)
| _ ->
- let (h,l) = decompose_app c in
- match kind_of_term h with
- | Const cst -> beta_applist (unfold cst,l)
+ let (h,l) = decompose_app_vect sigma c in
+ match EConstr.kind sigma h with
+ | Const cst -> beta_app sigma (unfold cst, l)
| _ -> c in
aux c
let betazetaevar_applist sigma n c l =
let rec stacklam n env t stack =
if Int.equal n 0 then applist (substl env t, stack) else
- match kind_of_term t, stack with
+ match EConstr.kind sigma t, stack with
| Lambda(_,_,c), arg::stacktl -> stacklam (n-1) (arg::env) c stacktl
| LetIn(_,b,_,c), _ -> stacklam (n-1) (substl env b::env) c stack
- | Evar ev, _ ->
- (match safe_evar_value sigma ev with
- | Some body -> stacklam n env body stack
- | None -> applist (substl env t, stack))
+ | Evar _, _ -> applist (substl env t, stack)
| _ -> anomaly (Pp.str "Not enough lambda/let's") in
stacklam n [] c l
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 4cd7a2a86..752c30a8a 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Univ
open Evd
open Environ
@@ -44,8 +45,8 @@ module Cst_stack : sig
val add_args : constr array -> t -> t
val add_cst : constr -> t -> t
val best_cst : t -> (constr * constr list) option
- val best_replace : constr -> t -> constr -> constr
- val reference : t -> Constant.t option
+ val best_replace : Evd.evar_map -> constr -> t -> constr -> constr
+ val reference : Evd.evar_map -> t -> Constant.t option
val pr : t -> Pp.std_ppcmds
end
@@ -63,7 +64,7 @@ module Stack : sig
| App of 'a app_node
| Case of case_info * 'a * 'a array * Cst_stack.t
| Proj of int * int * projection * Cst_stack.t
- | Fix of fixpoint * 'a t * Cst_stack.t
+ | Fix of ('a, 'a) pfixpoint * 'a t * Cst_stack.t
| Cst of cst_member * int (** current foccussed arg *) * int list (** remaining args *)
* 'a t * Cst_stack.t
| Shift of int
@@ -82,9 +83,9 @@ module Stack : sig
val compare_shape : 'a t -> 'a t -> bool
(** [fold2 f x sk1 sk2] folds [f] on any pair of term in [(sk1,sk2)].
@return the result and the lifts to apply on the terms *)
- val fold2 : ('a -> Term.constr -> Term.constr -> 'a) -> 'a ->
- Term.constr t -> Term.constr t -> 'a * int * int
- val map : (Term.constr -> Term.constr) -> Term.constr t -> Term.constr t
+ val fold2 : ('a -> constr -> constr -> 'a) -> 'a ->
+ constr t -> constr t -> 'a * int * int
+ val map : ('a -> 'a) -> 'a t -> 'a t
val append_app_list : 'a list -> 'a t -> 'a t
(** if [strip_app s] = [(a,b)], then [s = a @ b] and [b] does not
@@ -101,8 +102,8 @@ module Stack : sig
val tail : int -> 'a t -> 'a t
val nth : 'a t -> int -> 'a
- val best_state : constr * constr t -> Cst_stack.t -> constr * constr t
- val zip : ?refold:bool -> constr * constr t -> constr
+ val best_state : evar_map -> constr * constr t -> Cst_stack.t -> constr * constr t
+ val zip : ?refold:bool -> evar_map -> constr * constr t -> constr
end
(************************************************************************)
@@ -137,17 +138,18 @@ val strong_prodspine : local_reduction_function -> local_reduction_function
val stack_reduction_of_reduction :
'a reduction_function -> 'a state_reduction_function
i*)
-val stacklam : (state -> 'a) -> constr list -> constr -> constr Stack.t -> 'a
+val stacklam : (state -> 'a) -> constr list -> evar_map -> constr -> constr Stack.t -> 'a
val whd_state_gen : ?csts:Cst_stack.t -> refold:bool -> tactic_mode:bool ->
CClosure.RedFlags.reds -> Environ.env -> Evd.evar_map -> state -> state * Cst_stack.t
val iterate_whd_gen : bool -> CClosure.RedFlags.reds ->
- Environ.env -> Evd.evar_map -> Term.constr -> Term.constr
+ Environ.env -> Evd.evar_map -> constr -> constr
(** {6 Generic Optimized Reduction Function using Closures } *)
val clos_norm_flags : CClosure.RedFlags.reds -> reduction_function
+val clos_whd_flags : CClosure.RedFlags.reds -> reduction_function
(** Same as [(strong whd_beta[delta][iota])], but much faster on big terms *)
val nf_beta : local_reduction_function
@@ -200,9 +202,9 @@ val shrink_eta : constr -> constr
(** Various reduction functions *)
-val safe_evar_value : evar_map -> existential -> constr option
+val safe_evar_value : evar_map -> Constr.existential -> Constr.constr option
-val beta_applist : constr * constr list -> constr
+val beta_applist : evar_map -> constr * constr list -> constr
val hnf_prod_app : env -> evar_map -> constr -> constr -> constr
val hnf_prod_appvect : env -> evar_map -> constr -> constr array -> constr
@@ -213,12 +215,12 @@ val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
val splay_prod : env -> evar_map -> constr -> (Name.t * constr) list * constr
val splay_lam : env -> evar_map -> constr -> (Name.t * constr) list * constr
-val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * sorts
-val sort_of_arity : env -> evar_map -> constr -> sorts
-val splay_prod_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr
-val splay_lam_n : env -> evar_map -> int -> constr -> Context.Rel.t * constr
+val splay_arity : env -> evar_map -> constr -> (Name.t * constr) list * ESorts.t
+val sort_of_arity : env -> evar_map -> constr -> ESorts.t
+val splay_prod_n : env -> evar_map -> int -> constr -> rel_context * constr
+val splay_lam_n : env -> evar_map -> int -> constr -> rel_context * constr
val splay_prod_assum :
- env -> evar_map -> constr -> Context.Rel.t * constr
+ env -> evar_map -> constr -> rel_context * constr
type 'a miota_args = {
mP : constr; (** the result type *)
@@ -227,15 +229,15 @@ type 'a miota_args = {
mcargs : 'a list; (** the constructor's arguments *)
mlf : 'a array } (** the branch code vector *)
-val reducible_mind_case : constr -> bool
-val reduce_mind_case : constr miota_args -> constr
+val reducible_mind_case : evar_map -> constr -> bool
+val reduce_mind_case : evar_map -> constr miota_args -> constr
-val find_conclusion : env -> evar_map -> constr -> (constr,constr) kind_of_term
+val find_conclusion : env -> evar_map -> constr -> (constr, constr, ESorts.t, EInstance.t) kind_of_term
val is_arity : env -> evar_map -> constr -> bool
val is_sort : env -> evar_map -> types -> bool
-val contract_fix : ?env:Environ.env -> ?reference:Constant.t -> fixpoint -> constr
-val fix_recarg : fixpoint -> constr Stack.t -> (int * constr) option
+val contract_fix : ?env:Environ.env -> evar_map -> ?reference:Constant.t -> fixpoint -> constr
+val fix_recarg : ('a, 'a) pfixpoint -> 'b Stack.t -> (int * 'b) option
(** {6 Querying the kernel conversion oracle: opaque/transparent constants } *)
val is_transparent : Environ.env -> constant tableKey -> bool
@@ -274,14 +276,14 @@ val vm_infer_conv : ?pb:conv_pb -> env -> evar_map -> constr -> constr ->
(** [infer_conv_gen] behaves like [infer_conv] but is parametrized by a
conversion function. Used to pretype vm and native casts. *)
val infer_conv_gen : (conv_pb -> l2r:bool -> evar_map -> transparent_state ->
- (constr, evar_map) Reduction.generic_conversion_function) ->
+ (Constr.constr, evar_map) Reduction.generic_conversion_function) ->
?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state -> env ->
evar_map -> constr -> constr -> evar_map * bool
(** {6 Special-Purpose Reduction Functions } *)
-val whd_meta : evar_map -> constr -> constr
-val plain_instance : constr Metamap.t -> constr -> constr
+val whd_meta : local_reduction_function
+val plain_instance : evar_map -> constr Metamap.t -> constr -> constr
val instance : evar_map -> constr Metamap.t -> constr -> constr
val head_unfold_under_prod : transparent_state -> reduction_function
val betazetaevar_applist : evar_map -> int -> constr -> constr list -> constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 5b67af3e7..496c706ec 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -10,13 +10,14 @@ open Pp
open CErrors
open Util
open Term
-open Vars
open Inductive
open Inductiveops
open Names
open Reductionops
open Environ
open Termops
+open EConstr
+open Vars
open Arguments_renaming
open Context.Rel.Declaration
@@ -50,11 +51,11 @@ let anomaly_on_error f x =
with RetypeError e -> anomaly ~label:"retyping" (print_retype_error e)
let get_type_from_constraints env sigma t =
- if isEvar (fst (decompose_app t)) then
+ if isEvar sigma (fst (decompose_app_vect sigma t)) then
match
List.map_filter (fun (pbty,env,t1,t2) ->
- if is_fconv Reduction.CONV env sigma t t1 then Some t2
- else if is_fconv Reduction.CONV env sigma t t2 then Some t1
+ if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t1) then Some t2
+ else if is_fconv Reduction.CONV env sigma t (EConstr.of_constr t2) then Some t1
else None)
(snd (Evd.extract_all_conv_pbs sigma))
with
@@ -65,7 +66,7 @@ let get_type_from_constraints env sigma t =
let rec subst_type env sigma typ = function
| [] -> typ
| h::rest ->
- match kind_of_term (whd_all env sigma typ) with
+ match EConstr.kind sigma (whd_all env sigma typ) with
| Prod (na,c1,c2) -> subst_type env sigma (subst1 h c2) rest
| _ -> retype_error NonFunctionalConstruction
@@ -74,9 +75,9 @@ let rec subst_type env sigma typ = function
let sort_of_atomic_type env sigma ft args =
let rec concl_of_arity env n ar args =
- match kind_of_term (whd_all env sigma ar), args with
+ match EConstr.kind sigma (whd_all env sigma ar), args with
| Prod (na, t, b), h::l -> concl_of_arity (push_rel (LocalDef (na, lift n h, t)) env) (n + 1) b l
- | Sort s, [] -> s
+ | Sort s, [] -> ESorts.kind sigma s
| _ -> retype_error NotASort
in concl_of_arity env 0 ft (Array.to_list args)
@@ -85,37 +86,39 @@ let type_of_var env id =
with Not_found -> retype_error (BadVariable id)
let decomp_sort env sigma t =
- match kind_of_term (whd_all env sigma t) with
- | Sort s -> s
+ match EConstr.kind sigma (whd_all env sigma t) with
+ | Sort s -> ESorts.kind sigma s
| _ -> retype_error NotASort
+let destSort sigma s = ESorts.kind sigma (destSort sigma s)
+
let retype ?(polyprop=true) sigma =
let rec type_of env cstr =
- match kind_of_term cstr with
+ match EConstr.kind sigma cstr with
| Meta n ->
- (try strip_outer_cast (Evd.meta_ftype sigma n).Evd.rebus
+ (try strip_outer_cast sigma (EConstr.of_constr (Evd.meta_ftype sigma n).Evd.rebus)
with Not_found -> retype_error (BadMeta n))
| Rel n ->
let ty = RelDecl.get_type (lookup_rel n env) in
lift n ty
| Var id -> type_of_var env id
- | Const cst -> rename_type_of_constant env cst
- | Evar ev -> Evd.existential_type sigma ev
- | Ind ind -> rename_type_of_inductive env ind
- | Construct cstr -> rename_type_of_constructor env cstr
+ | Const (cst, u) -> EConstr.of_constr (rename_type_of_constant env (cst, EInstance.kind sigma u))
+ | Evar ev -> existential_type sigma ev
+ | Ind (ind, u) -> EConstr.of_constr (rename_type_of_inductive env (ind, EInstance.kind sigma u))
+ | Construct (cstr, u) -> EConstr.of_constr (rename_type_of_constructor env (cstr, EInstance.kind sigma u))
| Case (_,p,c,lf) ->
let Inductiveops.IndType(indf,realargs) =
let t = type_of env c in
try Inductiveops.find_rectype env sigma t
with Not_found ->
try
- let t = get_type_from_constraints env sigma t in
+ let t = EConstr.of_constr (get_type_from_constraints env sigma t) in
Inductiveops.find_rectype env sigma t
with Not_found -> retype_error BadRecursiveType
in
let n = inductive_nrealdecls_env env (fst (fst (dest_ind_family indf))) in
let t = betazetaevar_applist sigma n p realargs in
- (match kind_of_term (whd_all env sigma (type_of env t)) with
+ (match EConstr.kind sigma (whd_all env sigma (type_of env t)) with
| Prod _ -> whd_beta sigma (applist (t, [c]))
| _ -> t)
| Lambda (name,c1,c2) ->
@@ -124,25 +127,28 @@ let retype ?(polyprop=true) sigma =
subst1 b (type_of (push_rel (LocalDef (name,b,c1)) env) c2)
| Fix ((_,i),(_,tys,_)) -> tys.(i)
| CoFix (i,(_,tys,_)) -> tys.(i)
- | App(f,args) when is_template_polymorphic env f ->
+ | App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
- strip_outer_cast (subst_type env sigma t (Array.to_list args))
+ strip_outer_cast sigma (subst_type env sigma t (Array.to_list args))
| App(f,args) ->
- strip_outer_cast
+ strip_outer_cast sigma
(subst_type env sigma (type_of env f) (Array.to_list args))
| Proj (p,c) ->
let ty = type_of env c in
- (try
+ EConstr.of_constr (try
Inductiveops.type_of_projection_knowing_arg env sigma p c ty
with Invalid_argument _ -> retype_error BadRecursiveType)
| Cast (c,_, t) -> t
| Sort _ | Prod _ -> mkSort (sort_of env cstr)
and sort_of env t =
- match kind_of_term t with
- | Cast (c,_, s) when isSort s -> destSort s
- | Sort (Prop c) -> type1_sort
- | Sort (Type u) -> Type (Univ.super u)
+ match EConstr.kind sigma t with
+ | Cast (c,_, s) when isSort sigma s -> destSort sigma s
+ | Sort s ->
+ begin match ESorts.kind sigma s with
+ | Prop _ -> type1_sort
+ | Type u -> Type (Univ.super u)
+ end
| Prod (name,t,c2) ->
(match (sort_of env t, sort_of (push_rel (LocalAssum (name,t)) env) c2) with
| _, (Prop Null as s) -> s
@@ -152,7 +158,7 @@ let retype ?(polyprop=true) sigma =
| Prop Pos, (Type u2) -> Type (Univ.sup Univ.type0_univ u2)
| Prop Null, (Type _ as s) -> s
| Type u1, Type u2 -> Type (Univ.sup u1 u2))
- | App(f,args) when is_template_polymorphic env f ->
+ | App(f,args) when is_template_polymorphic env sigma f ->
let t = type_of_global_reference_knowing_parameters env f args in
sort_of_atomic_type env sigma t args
| App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args
@@ -160,15 +166,14 @@ let retype ?(polyprop=true) sigma =
| _ -> decomp_sort env sigma (type_of env t)
and sort_family_of env t =
- match kind_of_term t with
- | Cast (c,_, s) when isSort s -> family_of_sort (destSort s)
- | Sort (Prop c) -> InType
- | Sort (Type u) -> InType
+ match EConstr.kind sigma t with
+ | Cast (c,_, s) when isSort sigma s -> family_of_sort (destSort sigma s)
+ | Sort _ -> InType
| Prod (name,t,c2) ->
let s2 = sort_family_of (push_rel (LocalAssum (name,t)) env) c2 in
if not (is_impredicative_set env) &&
s2 == InSet && sort_family_of env t == InType then InType else s2
- | App(f,args) when is_template_polymorphic env f ->
+ | 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)
| App(f,args) ->
@@ -179,18 +184,22 @@ let retype ?(polyprop=true) sigma =
and type_of_global_reference_knowing_parameters env c args =
let argtyps =
- Array.map (fun c -> lazy (nf_evar sigma (type_of env c))) args in
- match kind_of_term c with
- | Ind ind ->
- let mip = lookup_mind_specif env (fst ind) in
- (try Inductive.type_of_inductive_knowing_parameters
- ~polyprop env (mip,snd ind) argtyps
+ Array.map (fun c -> lazy (EConstr.to_constr sigma (type_of env c))) args in
+ match EConstr.kind sigma c with
+ | Ind (ind, u) ->
+ let u = EInstance.kind sigma u in
+ let mip = lookup_mind_specif env ind in
+ EConstr.of_constr (try Inductive.type_of_inductive_knowing_parameters
+ ~polyprop env (mip, u) argtyps
with Reduction.NotArity -> retype_error NotAnArity)
- | Const cst ->
- (try Typeops.type_of_constant_knowing_parameters_in env cst argtyps
+ | Const (cst, u) ->
+ let u = EInstance.kind sigma u in
+ EConstr.of_constr (try Typeops.type_of_constant_knowing_parameters_in env (cst, u) argtyps
with Reduction.NotArity -> retype_error NotAnArity)
| Var id -> type_of_var env id
- | Construct cstr -> type_of_constructor env cstr
+ | Construct (cstr, u) ->
+ let u = EInstance.kind sigma u in
+ EConstr.of_constr (type_of_constructor env (cstr, u))
| _ -> assert false
in type_of, sort_of, sort_family_of,
@@ -204,17 +213,16 @@ let type_of_global_reference_knowing_parameters env sigma c args =
let _,_,_,f = retype sigma in anomaly_on_error (f env c) args
let type_of_global_reference_knowing_conclusion env sigma c conclty =
- let conclty = nf_evar sigma conclty in
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Ind (ind,u) ->
let spec = Inductive.lookup_mind_specif env ind in
- type_of_inductive_knowing_conclusion env sigma (spec,u) conclty
- | Const cst ->
- let t = constant_type_in env cst in
+ type_of_inductive_knowing_conclusion env sigma (spec, EInstance.kind sigma u) conclty
+ | Const (cst, u) ->
+ let t = constant_type_in env (cst, EInstance.kind sigma u) in
(* TODO *)
- sigma, Typeops.type_of_constant_type_knowing_parameters env t [||]
+ sigma, EConstr.of_constr (Typeops.type_of_constant_type_knowing_parameters env t [||])
| Var id -> sigma, type_of_var env id
- | Construct cstr -> sigma, type_of_constructor env cstr
+ | Construct (cstr, u) -> sigma, EConstr.of_constr (type_of_constructor env (cstr, EInstance.kind sigma u))
| _ -> assert false
(* Profiling *)
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index 8ca40f829..25129db1c 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -9,6 +9,7 @@
open Term
open Evd
open Environ
+open EConstr
(** This family of functions assumes its constr argument is known to be
well-typable. It does not type-check, just recompute the type
@@ -43,7 +44,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 -> Context.Rel.t -> sorts list
+val sorts_of_context : env -> evar_map -> rel_context -> sorts list
val expand_projection : env -> evar_map -> Names.projection -> constr -> constr list -> constr
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index 7da738508..67221046b 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -11,13 +11,14 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Libnames
open Globnames
open Termops
+open Environ
+open EConstr
+open Vars
open Find_subterm
open Namegen
-open Environ
open CClosure
open Reductionops
open Cbv
@@ -31,7 +32,7 @@ module NamedDecl = Context.Named.Declaration
(* Errors *)
type reduction_tactic_error =
- InvalidAbstraction of env * Evd.evar_map * constr * (env * Type_errors.type_error)
+ InvalidAbstraction of env * Evd.evar_map * EConstr.constr * (env * Type_errors.type_error)
exception ReductionTacticError of reduction_tactic_error
@@ -59,7 +60,8 @@ let is_evaluable env = function
let value_of_evaluable_ref env evref u =
match evref with
| EvalConstRef con ->
- (try constant_value_in env (con,u)
+ let u = Unsafe.to_instance u in
+ EConstr.of_constr (try constant_value_in env (con,u)
with NotEvaluableConst IsProj ->
raise (Invalid_argument "value_of_evaluable_ref"))
| EvalVarRef id -> env |> lookup_named id |> NamedDecl.get_value |> Option.get
@@ -77,14 +79,14 @@ type evaluable_reference =
| EvalConst of constant
| EvalVar of Id.t
| EvalRel of int
- | EvalEvar of existential
+ | EvalEvar of EConstr.existential
-let evaluable_reference_eq r1 r2 = match r1, r2 with
+let evaluable_reference_eq sigma r1 r2 = match r1, r2 with
| EvalConst c1, EvalConst c2 -> eq_constant c1 c2
| EvalVar id1, EvalVar id2 -> Id.equal id1 id2
| EvalRel i1, EvalRel i2 -> Int.equal i1 i2
| EvalEvar (e1, ctx1), EvalEvar (e2, ctx2) ->
- Evar.equal e1 e2 && Array.equal eq_constr ctx1 ctx2
+ Evar.equal e1 e2 && Array.equal (EConstr.eq_constr sigma) ctx1 ctx2
| _ -> false
let mkEvalRef ref u =
@@ -92,41 +94,49 @@ let mkEvalRef ref u =
| EvalConst cst -> mkConstU (cst,u)
| EvalVar id -> mkVar id
| EvalRel n -> mkRel n
- | EvalEvar ev -> mkEvar ev
+ | EvalEvar ev -> EConstr.mkEvar ev
-let isEvalRef env c = match kind_of_term c with
+let isEvalRef env sigma c = match EConstr.kind sigma c with
| Const (sp,_) -> is_evaluable env (EvalConstRef sp)
| Var id -> is_evaluable env (EvalVarRef id)
| Rel _ | Evar _ -> true
| _ -> false
-let destEvalRefU c = match kind_of_term c with
+let destEvalRefU sigma c = match EConstr.kind sigma c with
| Const (cst,u) -> EvalConst cst, u
- | Var id -> (EvalVar id, Univ.Instance.empty)
- | Rel n -> (EvalRel n, Univ.Instance.empty)
- | Evar ev -> (EvalEvar ev, Univ.Instance.empty)
+ | Var id -> (EvalVar id, EInstance.empty)
+ | Rel n -> (EvalRel n, EInstance.empty)
+ | Evar ev -> (EvalEvar ev, EInstance.empty)
| _ -> anomaly (Pp.str "Not an unfoldable reference")
let unsafe_reference_opt_value env sigma eval =
match eval with
| EvalConst cst ->
(match (lookup_constant cst env).Declarations.const_body with
- | Declarations.Def c -> Some (Mod_subst.force_constr c)
+ | Declarations.Def c -> Some (EConstr.of_constr (Mod_subst.force_constr c))
| _ -> None)
| EvalVar id ->
env |> lookup_named id |> NamedDecl.get_value
| EvalRel n ->
- env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value
- | EvalEvar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
+ | EvalEvar ev ->
+ match EConstr.kind sigma (mkEvar ev) with
+ | Evar _ -> None
+ | c -> Some (EConstr.of_kind c)
let reference_opt_value env sigma eval u =
match eval with
- | EvalConst cst -> constant_opt_value_in env (cst,u)
+ | EvalConst cst ->
+ let u = EInstance.kind sigma u in
+ Option.map EConstr.of_constr (constant_opt_value_in env (cst,u))
| EvalVar id ->
env |> lookup_named id |> NamedDecl.get_value
| EvalRel n ->
- env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value
- | EvalEvar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
+ | EvalEvar ev ->
+ match EConstr.kind sigma (mkEvar ev) with
+ | Evar _ -> None
+ | c -> Some (EConstr.of_kind c)
exception NotEvaluable
let reference_value env sigma c u =
@@ -177,18 +187,18 @@ let eval_table = Summary.ref (Cmap.empty : frozen) ~name:"evaluation"
the xp..x1.
*)
-let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
+let check_fix_reversibility sigma labs args ((lv,i),(_,tys,bds)) =
let n = List.length labs in
let nargs = List.length args in
if nargs > n then raise Elimconst;
let nbfix = Array.length bds in
let li =
List.map
- (function d -> match kind_of_term d with
+ (function d -> match EConstr.kind sigma d with
| Rel k ->
if
- Array.for_all (noccurn k) tys
- && Array.for_all (noccurn (k+nbfix)) bds
+ Array.for_all (Vars.noccurn sigma k) tys
+ && Array.for_all (Vars.noccurn sigma (k+nbfix)) bds
&& k <= n
then
(k, List.nth labs (k-1))
@@ -202,7 +212,7 @@ let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
raise Elimconst;
List.iteri (fun i t_i ->
if not (Int.List.mem_assoc (i+1) li) then
- let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels t_i)) in
+ let fvs = List.map ((+) (i+1)) (Int.Set.elements (free_rels sigma t_i)) in
match List.intersect Int.equal fvs reversible_rels with
| [] -> ()
| _ -> raise Elimconst)
@@ -238,10 +248,11 @@ let invert_name labs l na0 env sigma ref = function
try match unsafe_reference_opt_value env sigma ref with
| None -> None
| Some c ->
- let labs',ccl = decompose_lam c in
+ let labs',ccl = decompose_lam sigma c in
let _, l' = whd_betalet_stack sigma ccl in
let labs' = List.map snd labs' in
(** ppedrot: there used to be generic equality on terms here *)
+ let eq_constr c1 c2 = EConstr.eq_constr sigma c1 c2 in
if List.equal eq_constr labs' labs &&
List.equal eq_constr l l' then Some (minfxargs,ref)
else None
@@ -256,16 +267,16 @@ let invert_name labs l na0 env sigma ref = function
let compute_consteval_direct env sigma ref =
let rec srec env n labs onlyproj c =
let c',l = whd_betadeltazeta_stack env sigma c in
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Lambda (id,t,g) when List.is_empty l && not onlyproj ->
let open Context.Rel.Declaration in
srec (push_rel (LocalAssum (id,t)) env) (n+1) (t::labs) onlyproj g
| Fix fix when not onlyproj ->
- (try check_fix_reversibility labs l fix
+ (try check_fix_reversibility sigma labs l fix
with Elimconst -> NotAnElimination)
- | Case (_,_,d,_) when isRel d && not onlyproj -> EliminationCases n
+ | Case (_,_,d,_) when isRel sigma d && not onlyproj -> EliminationCases n
| Case (_,_,d,_) -> srec env n labs true d
- | Proj (p, d) when isRel d -> EliminationProj n
+ | Proj (p, d) when isRel sigma d -> EliminationProj n
| _ -> NotAnElimination
in
match unsafe_reference_opt_value env sigma ref with
@@ -276,7 +287,7 @@ let compute_consteval_mutual_fix env sigma ref =
let rec srec env minarg labs ref c =
let c',l = whd_betalet_stack sigma c in
let nargs = List.length l in
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Lambda (na,t,g) when List.is_empty l ->
let open Context.Rel.Declaration in
srec (push_rel (LocalAssum (na,t)) env) (minarg+1) (t::labs) ref g
@@ -292,9 +303,9 @@ let compute_consteval_mutual_fix env sigma ref =
let new_minarg = max (minarg'+minarg-nargs) minarg' in
EliminationMutualFix (new_minarg,ref,(refs,infos))
| _ -> assert false)
- | _ when isEvalRef env c' ->
+ | _ when isEvalRef env sigma c' ->
(* Forget all \'s and args and do as if we had started with c' *)
- let ref,_ = destEvalRefU c' in
+ let ref,_ = destEvalRefU sigma c' in
(match unsafe_reference_opt_value env sigma ref with
| None -> anomaly (Pp.str "Should have been trapped by compute_direct")
| Some c -> srec env (minarg-nargs) [] ref c)
@@ -353,17 +364,17 @@ let make_elim_fun (names,(nbfix,lv,n)) u largs =
(* k from the comment is q+1 *)
try mkRel (p+1-(List.index Int.equal (n-q) lyi))
with Not_found -> aq)
- 0 (List.map (lift p) lu)
+ 0 (List.map (Vars.lift p) lu)
in
fun i ->
match names.(i) with
| None -> None
| Some (minargs,ref) ->
- let body = applistc (mkEvalRef ref u) la in
+ let body = applist (mkEvalRef ref u, la) in
let g =
List.fold_left_i (fun q (* j = n+1-q *) c (ij,tij) ->
- let subst = List.map (lift (-q)) (List.firstn (n-ij) la) in
- let tij' = substl (List.rev subst) tij in
+ let subst = List.map (Vars.lift (-q)) (List.firstn (n-ij) la) in
+ let tij' = Vars.substl (List.rev subst) tij in
mkLambda (x,tij',c)) 1 body (List.rev lv)
in Some (minargs,g)
@@ -383,7 +394,7 @@ let substl_with_function subst sigma constr =
let evd = ref sigma in
let minargs = ref Evar.Map.empty in
let v = Array.of_list subst in
- let rec subst_total k c = match kind_of_term c with
+ let rec subst_total k c = match EConstr.kind sigma c with
| Rel i when k < i ->
if i <= k + Array.length v then
match v.(i-k-1) with
@@ -393,11 +404,11 @@ let substl_with_function subst sigma constr =
let sigma = Sigma.to_evar_map sigma in
evd := sigma;
minargs := Evar.Map.add evk min !minargs;
- lift k (mkEvar (evk, [|fx;ref|]))
- | (fx, None) -> lift k fx
+ Vars.lift k (mkEvar (evk, [|fx;ref|]))
+ | (fx, None) -> Vars.lift k fx
else mkRel (i - Array.length v)
| _ ->
- map_constr_with_binders succ subst_total k c in
+ map_with_binders sigma succ subst_total k c in
let c = subst_total 0 constr in
(c, !evd, !minargs)
@@ -407,28 +418,28 @@ exception Partial
reduction is solved by the expanded fix term. *)
let solve_arity_problem env sigma fxminargs c =
let evm = ref sigma in
- let set_fix i = evm := Evd.define i (mkVar vfx) !evm in
+ let set_fix i = evm := Evd.define i (Constr.mkVar vfx) !evm in
let rec check strict c =
let c' = whd_betaiotazeta sigma c in
- let (h,rcargs) = decompose_app c' in
- match kind_of_term h with
+ let (h,rcargs) = decompose_app_vect sigma c' in
+ match EConstr.kind sigma h with
Evar(i,_) when Evar.Map.mem i fxminargs && not (Evd.is_defined !evm i) ->
let minargs = Evar.Map.find i fxminargs in
- if List.length rcargs < minargs then
+ if Array.length rcargs < minargs then
if strict then set_fix i
else raise Partial;
- List.iter (check strict) rcargs
- | (Var _|Const _) when isEvalRef env h ->
- (let ev, u = destEvalRefU h in
+ Array.iter (check strict) rcargs
+ | (Var _|Const _) when isEvalRef env sigma h ->
+ (let ev, u = destEvalRefU sigma h in
match reference_opt_value env sigma ev u with
| Some h' ->
let bak = !evm in
- (try List.iter (check false) rcargs
+ (try Array.iter (check false) rcargs
with Partial ->
evm := bak;
- check strict (applist(h',rcargs)))
- | None -> List.iter (check strict) rcargs)
- | _ -> iter_constr (check strict) c' in
+ check strict (mkApp(h',rcargs)))
+ | None -> Array.iter (check strict) rcargs)
+ | _ -> EConstr.iter sigma (check strict) c' in
check true c;
!evm
@@ -439,16 +450,18 @@ let substl_checking_arity env subst sigma c =
let sigma' = solve_arity_problem env sigma minargs body in
(* we propagate the constraints: solved problems are substituted;
the other ones are replaced by the function symbol *)
- let rec nf_fix c =
- match kind_of_term c with
- Evar(i,[|fx;f|] as ev) when Evar.Map.mem i minargs ->
- (match Evd.existential_opt_value sigma' ev with
- Some c' -> c'
- | None -> f)
- | _ -> map_constr nf_fix c in
+ let rec nf_fix c = match EConstr.kind sigma c with
+ | Evar (i,[|fx;f|]) when Evar.Map.mem i minargs ->
+ (** FIXME: find a less hackish way of doing this *)
+ begin match EConstr.kind sigma' c with
+ | Evar _ -> f
+ | c -> EConstr.of_kind c
+ end
+ | _ -> EConstr.map sigma nf_fix c
+ in
nf_fix body
-type fix_reduction_result = NotReducible | Reduced of (constr*constr list)
+type fix_reduction_result = NotReducible | Reduced of (constr * constr list)
let reduce_fix whdfun sigma fix stack =
match fix_recarg fix (Stack.append_app_list stack Stack.empty) with
@@ -456,8 +469,8 @@ let reduce_fix whdfun sigma fix stack =
| Some (recargnum,recarg) ->
let (recarg'hd,_ as recarg') = whdfun sigma recarg in
let stack' = List.assign stack recargnum (applist recarg') in
- (match kind_of_term recarg'hd with
- | Construct _ -> Reduced (contract_fix fix, stack')
+ (match EConstr.kind sigma recarg'hd with
+ | Construct _ -> Reduced (contract_fix sigma fix, stack')
| _ -> NotReducible)
let contract_fix_use_function env sigma f
@@ -472,13 +485,13 @@ let reduce_fix_use_function env sigma f whfun fix stack =
| None -> NotReducible
| Some (recargnum,recarg) ->
let (recarg'hd,_ as recarg') =
- if isRel recarg then
+ if EConstr.isRel sigma recarg then
(* The recarg cannot be a local def, no worry about the right env *)
(recarg, [])
else
whfun recarg in
let stack' = List.assign stack recargnum (applist recarg') in
- (match kind_of_term recarg'hd with
+ (match EConstr.kind sigma recarg'hd with
| Construct _ ->
Reduced (contract_fix_use_function env sigma f fix,stack')
| _ -> NotReducible)
@@ -492,13 +505,13 @@ let contract_cofix_use_function env sigma f
sigma (nf_beta sigma bodies.(bodynum))
let reduce_mind_case_use_function func env sigma mia =
- match kind_of_term mia.mconstr with
+ match EConstr.kind sigma mia.mconstr with
| Construct ((ind_sp,i),u) ->
let real_cargs = List.skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1), real_cargs)
| CoFix (bodynum,(names,_,_) as cofix) ->
let build_cofix_name =
- if isConst func then
+ if isConst sigma func then
let minargs = List.length mia.mcargs in
fun i ->
if Int.equal i bodynum then Some (minargs,func)
@@ -509,13 +522,13 @@ let reduce_mind_case_use_function func env sigma mia =
mutual inductive, try to reuse the global name if
the block was indeed initially built as a global
definition *)
- let kn = map_puniverses (fun x -> con_with_label x (Label.of_id id))
- (destConst func)
- in
- try match constant_opt_value_in env kn with
+ let (kn, u) = destConst sigma func in
+ let kn = con_with_label kn (Label.of_id id) in
+ let cst = (kn, EInstance.kind sigma u) in
+ try match constant_opt_value_in env cst with
| None -> None
(* TODO: check kn is correct *)
- | Some _ -> Some (minargs,mkConstU kn)
+ | Some _ -> Some (minargs,mkConstU (kn, u))
with Not_found -> None
else
fun _ -> None in
@@ -525,43 +538,43 @@ let reduce_mind_case_use_function func env sigma mia =
| _ -> assert false
-let match_eval_ref env constr =
- match kind_of_term constr with
+let match_eval_ref env sigma constr =
+ match EConstr.kind sigma constr with
| Const (sp, u) when is_evaluable env (EvalConstRef sp) ->
Some (EvalConst sp, u)
- | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, Univ.Instance.empty)
- | Rel i -> Some (EvalRel i, Univ.Instance.empty)
- | Evar ev -> Some (EvalEvar ev, Univ.Instance.empty)
+ | Var id when is_evaluable env (EvalVarRef id) -> Some (EvalVar id, EInstance.empty)
+ | Rel i -> Some (EvalRel i, EInstance.empty)
+ | Evar ev -> Some (EvalEvar ev, EInstance.empty)
| _ -> None
let match_eval_ref_value env sigma constr =
- match kind_of_term constr with
+ match EConstr.kind sigma constr with
| Const (sp, u) when is_evaluable env (EvalConstRef sp) ->
- Some (constant_value_in env (sp, u))
+ let u = EInstance.kind sigma u in
+ Some (EConstr.of_constr (constant_value_in env (sp, u)))
| Var id when is_evaluable env (EvalVarRef id) ->
env |> lookup_named id |> NamedDecl.get_value
| Rel n ->
- env |> lookup_rel n |> RelDecl.map_value (lift n) |> RelDecl.get_value
- | Evar ev -> Evd.existential_opt_value sigma ev
+ env |> lookup_rel n |> RelDecl.get_value |> Option.map (lift n)
| _ -> None
let special_red_case env sigma whfun (ci, p, c, lf) =
let rec redrec s =
let (constr, cargs) = whfun s in
- match match_eval_ref env constr with
+ match match_eval_ref env sigma constr with
| Some (ref, u) ->
(match reference_opt_value env sigma ref u with
| None -> raise Redelimination
| Some gvalue ->
- if reducible_mind_case gvalue then
+ if reducible_mind_case sigma gvalue then
reduce_mind_case_use_function constr env sigma
{mP=p; mconstr=gvalue; mcargs=cargs;
mci=ci; mlf=lf}
else
redrec (applist(gvalue, cargs)))
| None ->
- if reducible_mind_case constr then
- reduce_mind_case
+ if reducible_mind_case sigma constr then
+ reduce_mind_case sigma
{mP=p; mconstr=constr; mcargs=cargs;
mci=ci; mlf=lf}
else
@@ -574,7 +587,7 @@ let recargs = function
| EvalConst c -> ReductionBehaviour.get (ConstRef c)
let reduce_projection env sigma pb (recarg'hd,stack') stack =
- (match kind_of_term recarg'hd with
+ (match EConstr.kind sigma recarg'hd with
| Construct _ ->
let proj_narg =
pb.Declarations.proj_npars + pb.Declarations.proj_arg
@@ -583,11 +596,11 @@ let reduce_projection env sigma pb (recarg'hd,stack') stack =
let reduce_proj env sigma whfun whfun' c =
let rec redrec s =
- match kind_of_term s with
+ match EConstr.kind sigma s with
| Proj (proj, c) ->
let c' = try redrec c with Redelimination -> c in
let constr, cargs = whfun c' in
- (match kind_of_term constr with
+ (match EConstr.kind sigma constr with
| Construct _ ->
let proj_narg =
let pb = lookup_projection proj env in
@@ -604,7 +617,7 @@ let reduce_proj env sigma whfun whfun' c =
let whd_nothing_for_iota env sigma s =
let rec whrec (x, stack as s) =
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Rel n ->
let open Context.Rel.Declaration in
(match lookup_rel n env with
@@ -615,27 +628,26 @@ let whd_nothing_for_iota env sigma s =
(match lookup_named id env with
| LocalDef (_,body,_) -> whrec (body, stack)
| _ -> s)
- | Evar ev ->
- (try whrec (Evd.existential_value sigma ev, stack)
- with Evd.NotInstantiatedEvar | Not_found -> s)
+ | Evar ev -> s
| Meta ev ->
- (try whrec (Evd.meta_value sigma ev, stack)
+ (try whrec (EConstr.of_constr (Evd.meta_value sigma ev), stack)
with Not_found -> s)
- | Const const when is_transparent_constant full_transparent_state (fst const) ->
- (match constant_opt_value_in env const with
- | Some body -> whrec (body, stack)
+ | Const (const, u) when is_transparent_constant full_transparent_state const ->
+ let u = EInstance.kind sigma u in
+ (match constant_opt_value_in env (const, u) with
+ | Some body -> whrec (EConstr.of_constr body, stack)
| None -> s)
- | LetIn (_,b,_,c) -> stacklam whrec [b] c stack
+ | LetIn (_,b,_,c) -> stacklam whrec [b] sigma c stack
| Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, Stack.append_app cl stack)
| Lambda (na,t,c) ->
(match Stack.decomp stack with
- | Some (a,m) -> stacklam whrec [a] c m
+ | Some (a,m) -> stacklam whrec [a] sigma c m
| _ -> s)
| x -> s
in
- decompose_app (Stack.zip (whrec (s,Stack.empty)))
+ EConstr.decompose_app sigma (Stack.zip sigma (whrec (s,Stack.empty)))
(* [red_elim_const] contracts iota/fix/cofix redexes hidden behind
constants by keeping the name of the constants in the recursive calls;
@@ -660,7 +672,7 @@ let rec red_elim_const env sigma ref u largs =
let c = reference_value env sigma ref u in
let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
let whfun = whd_simpl_stack env sigma in
- (special_red_case env sigma whfun (destCase c'), lrest), nocase
+ (special_red_case env sigma whfun (EConstr.destCase sigma c'), lrest), nocase
| EliminationProj n when nargs >= n ->
let c = reference_value env sigma ref u in
let c', lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
@@ -672,22 +684,22 @@ let rec red_elim_const env sigma ref u largs =
let d, lrest = whd_nothing_for_iota env sigma (applist(c,largs)) in
let f = make_elim_fun ([|Some (minfxargs,ref)|],infos) u largs in
let whfun = whd_construct_stack env sigma in
- (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with
+ (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
| Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
| EliminationMutualFix (min,refgoal,refinfos) when nargs >= min ->
let rec descend (ref,u) args =
let c = reference_value env sigma ref u in
- if evaluable_reference_eq ref refgoal then
+ if evaluable_reference_eq sigma ref refgoal then
(c,args)
else
let c', lrest = whd_betalet_stack sigma (applist(c,args)) in
- descend (destEvalRefU c') lrest in
+ descend (destEvalRefU sigma c') lrest in
let (_, midargs as s) = descend (ref,u) largs in
let d, lrest = whd_nothing_for_iota env sigma (applist s) in
let f = make_elim_fun refinfos u midargs in
let whfun = whd_construct_stack env sigma in
- (match reduce_fix_use_function env sigma f whfun (destFix d) lrest with
+ (match reduce_fix_use_function env sigma f whfun (destFix sigma d) lrest with
| NotReducible -> raise Redelimination
| Reduced (c,rest) -> (nf_beta sigma c, rest), nocase)
| NotAnElimination when unfold_nonelim ->
@@ -705,7 +717,7 @@ and reduce_params env sigma stack l =
else
let arg = List.nth stack i in
let rarg = whd_construct_stack env sigma arg in
- match kind_of_term (fst rarg) with
+ match EConstr.kind sigma (fst rarg) with
| Construct _ -> List.assign stack i (applist rarg)
| _ -> raise Redelimination)
stack l
@@ -716,13 +728,15 @@ and reduce_params env sigma stack l =
and whd_simpl_stack env sigma =
let rec redrec s =
- let (x, stack as s') = decompose_app s in
- match kind_of_term x with
+ let (x, stack) = decompose_app_vect sigma s in
+ let stack = Array.to_list stack in
+ let s' = (x, stack) in
+ match EConstr.kind sigma x with
| Lambda (na,t,c) ->
(match stack with
| [] -> s'
- | a :: rest -> redrec (beta_applist (x,stack)))
- | LetIn (n,b,t,c) -> redrec (applist (substl [b] c, stack))
+ | a :: rest -> redrec (beta_applist sigma (x, stack)))
+ | LetIn (n,b,t,c) -> redrec (applist (Vars.substl [b] c, stack))
| App (f,cl) -> redrec (applist(f, (Array.to_list cl)@stack))
| Cast (c,_,_) -> redrec (applist(c, stack))
| Case (ci,p,c,lf) ->
@@ -762,12 +776,12 @@ and whd_simpl_stack env sigma =
with Redelimination -> s')
| _ ->
- match match_eval_ref env x with
+ match match_eval_ref env sigma x with
| Some (ref, u) ->
(try
let sapp, nocase = red_elim_const env sigma ref u stack in
let hd, _ as s'' = redrec (applist(sapp)) in
- let rec is_case x = match kind_of_term x with
+ let rec is_case x = match EConstr.kind sigma x with
| Lambda (_,_, x) | LetIn (_,_,_, x) | Cast (x, _,_) -> is_case x
| App (hd, _) -> is_case hd
| Case _ -> true
@@ -783,8 +797,8 @@ and whd_simpl_stack env sigma =
and whd_construct_stack env sigma s =
let (constr, cargs as s') = whd_simpl_stack env sigma s in
- if reducible_mind_case constr then s'
- else match match_eval_ref env constr with
+ if reducible_mind_case sigma constr then s'
+ else match match_eval_ref env sigma constr with
| Some (ref, u) ->
(match reference_opt_value env sigma ref u with
| None -> raise Redelimination
@@ -800,12 +814,12 @@ and whd_construct_stack env sigma s =
*)
let try_red_product env sigma c =
- let simpfun = clos_norm_flags betaiotazeta env sigma in
+ let simpfun c = clos_norm_flags betaiotazeta env sigma c in
let rec redrec env x =
let x = whd_betaiota sigma x in
- match kind_of_term x with
+ match EConstr.kind sigma x with
| App (f,l) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Fix fix ->
let stack = Stack.append_app l Stack.empty in
(match fix_recarg fix stack with
@@ -813,17 +827,17 @@ let try_red_product env sigma c =
| Some (recargnum,recarg) ->
let recarg' = redrec env recarg in
let stack' = Stack.assign stack recargnum recarg' in
- simpfun (Stack.zip (f,stack')))
- | _ -> simpfun (appvect (redrec env f, l)))
+ simpfun (Stack.zip sigma (f,stack')))
+ | _ -> simpfun (mkApp (redrec env f, l)))
| Cast (c,_,_) -> redrec env c
| Prod (x,a,b) ->
let open Context.Rel.Declaration in
- mkProd (x, a, redrec (push_rel (LocalAssum (x,a)) env) b)
- | LetIn (x,a,b,t) -> redrec env (subst1 a t)
+ mkProd (x, a, redrec (push_rel (LocalAssum (x, a)) env) b)
+ | LetIn (x,a,b,t) -> redrec env (Vars.subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
| Proj (p, c) ->
let c' =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Construct _ -> c
| _ -> redrec env c
in
@@ -832,7 +846,7 @@ let try_red_product env sigma c =
| Reduced s -> simpfun (applist s)
| NotReducible -> raise Redelimination)
| _ ->
- (match match_eval_ref env x with
+ (match match_eval_ref env sigma x with
| Some (ref, u) ->
(* TO DO: re-fold fixpoints after expansion *)
(* to get true one-step reductions *)
@@ -915,10 +929,10 @@ let whd_simpl_orelse_delta_but_fix env sigma c =
let (constr, stack as s') = whd_simpl_stack env sigma s in
match match_eval_ref_value env sigma constr with
| Some c ->
- (match kind_of_term (strip_lam c) with
+ (match EConstr.kind sigma (snd (decompose_lam sigma c)) with
| CoFix _ | Fix _ -> s'
| Proj (p,t) when
- (match kind_of_term constr with
+ (match EConstr.kind sigma constr with
| Const (c', _) -> eq_constant (Projection.constant p) c'
| _ -> false) ->
let pb = Environ.lookup_projection p env in
@@ -944,9 +958,9 @@ let simpl env sigma c = strong whd_simpl env sigma c
(* Reduction at specific subterms *)
let matches_head env sigma c t =
- match kind_of_term t with
+ match EConstr.kind sigma t with
| App (f,_) -> Constr_matching.matches env sigma c f
- | Proj (p, _) -> Constr_matching.matches env sigma c (mkConst (Projection.constant p))
+ | Proj (p, _) -> Constr_matching.matches env sigma c (mkConstU (Projection.constant p, EInstance.empty))
| _ -> raise Constr_matching.PatternMatchingFailure
(** FIXME: Specific function to handle projections: it ignores what happens on the
@@ -954,20 +968,20 @@ let matches_head env sigma c t =
of the projection and its eta expanded form.
*)
let change_map_constr_with_binders_left_to_right g f (env, l as acc) sigma c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, r) -> (* Treat specially for partial applications *)
let t = Retyping.expand_projection env sigma p r [] in
- let hdf, al = destApp t in
+ let hdf, al = destApp sigma t in
let a = al.(Array.length al - 1) in
let app = (mkApp (hdf, Array.sub al 0 (Array.length al - 1))) in
let app' = f acc app in
let a' = f acc a in
- (match kind_of_term app' with
+ (match EConstr.kind sigma app' with
| App (hdf', al') when hdf' == hdf ->
(* Still the same projection, we ignore the change in parameters *)
mkProj (p, a')
| _ -> mkApp (app', [| a' |]))
- | _ -> map_constr_with_binders_left_to_right g f acc c
+ | _ -> map_constr_with_binders_left_to_right sigma g f acc c
let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
let (nowhere_except_in,locs) = Locusops.convert_occs occs in
@@ -1003,7 +1017,7 @@ let e_contextually byhead (occs,c) f = { e_redfun = begin fun env sigma t ->
and traverse_below nested envc t =
(* when byhead, find other occurrences without matching again partial
application with same head *)
- match kind_of_term t with
+ match EConstr.kind !evd t with
| App (f,l) when byhead -> mkApp (f, Array.map_left (traverse nested envc) l)
| Proj (p,c) when byhead -> mkProj (p,traverse nested envc c)
| _ ->
@@ -1028,9 +1042,9 @@ let contextually byhead occs f env sigma t =
* ol is the occurrence list to find. *)
let match_constr_evaluable_ref sigma c evref =
- match kind_of_term c, evref with
+ match EConstr.kind sigma c, evref with
| Const (c,u), EvalConstRef c' when eq_constant c c' -> Some u
- | Var id, EvalVarRef id' when id_eq id id' -> Some Univ.Instance.empty
+ | Var id, EvalVarRef id' when id_eq id id' -> Some EInstance.empty
| _, _ -> None
let substlin env sigma evalref n (nowhere_except_in,locs) c =
@@ -1049,7 +1063,7 @@ let substlin env sigma evalref n (nowhere_except_in,locs) c =
incr pos;
if ok then value u else c
| None ->
- map_constr_with_binders_left_to_right
+ map_constr_with_binders_left_to_right sigma
(fun _ () -> ())
substrec () c
in
@@ -1062,9 +1076,9 @@ let string_of_evaluable_ref env = function
string_of_qualid
(Nametab.shortest_qualid_of_global (vars_of_env env) (ConstRef kn))
-let unfold env sigma name =
+let unfold env sigma name c =
if is_evaluable env name then
- clos_norm_flags (unfold_red name) env sigma
+ clos_norm_flags (unfold_red name) env sigma c
else
error (string_of_evaluable_ref env name^" is opaque.")
@@ -1102,17 +1116,17 @@ let fold_one_com com env sigma c =
(* Reason first on the beta-iota-zeta normal form of the constant as
unfold produces it, so that the "unfold f; fold f" configuration works
to refold fix expressions *)
- let a = subst_term (clos_norm_flags unfold_side_red env sigma rcom) c in
- if not (eq_constr a c) then
- subst1 com a
+ let a = subst_term sigma (clos_norm_flags unfold_side_red env sigma rcom) c in
+ if not (EConstr.eq_constr sigma a c) then
+ Vars.subst1 com a
else
(* Then reason on the non beta-iota-zeta form for compatibility -
even if it is probably a useless configuration *)
- let a = subst_term rcom c in
- subst1 com a
+ let a = subst_term sigma rcom c in
+ Vars.subst1 com a
let fold_commands cl env sigma c =
- List.fold_right (fun com -> fold_one_com com env sigma) (List.rev cl) c
+ List.fold_right (fun com c -> fold_one_com com env sigma c) (List.rev cl) c
(* call by value reduction functions *)
@@ -1130,11 +1144,11 @@ let compute = cbv_betadeltaiota
(* gives [na:ta]c' such that c converts to ([na:ta]c' a), abstracting only
* the specified occurrences. *)
-let abstract_scheme env (locc,a) (c, sigma) =
+let abstract_scheme env sigma (locc,a) (c, sigma) =
let ta = Retyping.get_type_of env sigma a in
- let na = named_hd env ta Anonymous in
- if occur_meta ta then error "Cannot find a type for the generalisation.";
- if occur_meta a then
+ let na = named_hd env sigma ta Anonymous in
+ if occur_meta sigma ta then error "Cannot find a type for the generalisation.";
+ if occur_meta sigma a then
mkLambda (na,ta,c), sigma
else
let c', sigma' = subst_closed_term_occ env sigma (AtOccs locc) a c in
@@ -1142,7 +1156,7 @@ let abstract_scheme env (locc,a) (c, sigma) =
let pattern_occs loccs_trm = { e_redfun = begin fun env sigma c ->
let sigma = Sigma.to_evar_map sigma in
- let abstr_trm, sigma = List.fold_right (abstract_scheme env) loccs_trm (c,sigma) in
+ let abstr_trm, sigma = List.fold_right (abstract_scheme env sigma) loccs_trm (c,sigma) in
try
let _ = Typing.unsafe_type_of env sigma abstr_trm in
Sigma.Unsafe.of_pair (applist(abstr_trm, List.map snd loccs_trm), sigma)
@@ -1171,7 +1185,7 @@ let check_not_primitive_record env ind =
let reduce_to_ind_gen allow_product env sigma t =
let rec elimrec env t l =
let t = hnf_constr env sigma t in
- match kind_of_term (fst (decompose_app t)) with
+ match EConstr.kind sigma (fst (decompose_app_vect sigma t)) with
| Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t l)
| Prod (n,ty,t') ->
let open Context.Rel.Declaration in
@@ -1183,18 +1197,18 @@ let reduce_to_ind_gen allow_product env sigma t =
(* Last chance: we allow to bypass the Opaque flag (as it
was partially the case between V5.10 and V8.1 *)
let t' = whd_all env sigma t in
- match kind_of_term (fst (decompose_app t')) with
+ match EConstr.kind sigma (fst (decompose_app_vect sigma t')) with
| Ind ind-> (check_privacy env ind, it_mkProd_or_LetIn t' l)
| _ -> user_err (str"Not an inductive product.")
in
elimrec env t []
-let reduce_to_quantified_ind x = reduce_to_ind_gen true x
-let reduce_to_atomic_ind x = reduce_to_ind_gen false x
+let reduce_to_quantified_ind env sigma c = reduce_to_ind_gen true env sigma c
+let reduce_to_atomic_ind env sigma c = reduce_to_ind_gen false env sigma c
let find_hnf_rectype env sigma t =
let ind,t = reduce_to_atomic_ind env sigma t in
- ind, snd (decompose_app t)
+ ind, snd (decompose_app sigma t)
(* Reduce the weak-head redex [beta,iota/fix/cofix[all],cast,zeta,simpl/delta]
or raise [NotStepReducible] if not a weak-head redex *)
@@ -1203,13 +1217,13 @@ exception NotStepReducible
let one_step_reduce env sigma c =
let rec redrec (x, stack) =
- match kind_of_term x with
+ match EConstr.kind sigma x with
| Lambda (n,t,c) ->
(match stack with
| [] -> raise NotStepReducible
- | a :: rest -> (subst1 a c, rest))
+ | a :: rest -> (Vars.subst1 a c, rest))
| App (f,cl) -> redrec (f, (Array.to_list cl)@stack)
- | LetIn (_,f,_,cl) -> (subst1 f cl,stack)
+ | LetIn (_,f,_,cl) -> (Vars.subst1 f cl,stack)
| Cast (c,_,_) -> redrec (c,stack)
| Case (ci,p,c,lf) ->
(try
@@ -1221,8 +1235,8 @@ let one_step_reduce env sigma c =
| Reduced s' -> s'
| NotReducible -> raise NotStepReducible
with Redelimination -> raise NotStepReducible)
- | _ when isEvalRef env x ->
- let ref,u = destEvalRefU x in
+ | _ when isEvalRef env sigma x ->
+ let ref,u = destEvalRefU sigma x in
(try
fst (red_elim_const env sigma ref u stack)
with Redelimination ->
@@ -1249,8 +1263,8 @@ let reduce_to_ref_gen allow_product env sigma ref t =
else
(* lazily reduces to match the head of [t] with the expected [ref] *)
let rec elimrec env t l =
- let c, _ = decompose_appvect (Reductionops.whd_nored sigma t) in
- match kind_of_term c with
+ let c, _ = decompose_app_vect sigma t in
+ match EConstr.kind sigma c with
| Prod (n,ty,t') ->
if allow_product then
let open Context.Rel.Declaration in
@@ -1259,7 +1273,7 @@ let reduce_to_ref_gen allow_product env sigma ref t =
error_cannot_recognize ref
| _ ->
try
- if eq_gr (global_of_constr c) ref
+ if eq_gr (fst (global_of_constr sigma c)) ref
then it_mkProd_or_LetIn t l
else raise Not_found
with Not_found ->
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index f8dfe1adf..76d0bc241 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -10,6 +10,7 @@ open Names
open Term
open Environ
open Evd
+open EConstr
open Reductionops
open Pattern
open Globnames
@@ -75,12 +76,12 @@ val cbv_norm_flags : CClosure.RedFlags.reds -> reduction_function
(** [reduce_to_atomic_ind env sigma t] puts [t] in the form [t'=(I args)]
with [I] an inductive definition;
returns [I] and [t'] or fails with a user error *)
-val reduce_to_atomic_ind : env -> evar_map -> types -> pinductive * types
+val reduce_to_atomic_ind : env -> evar_map -> types -> (inductive * EInstance.t) * types
(** [reduce_to_quantified_ind env sigma t] puts [t] in the form
[t'=(x1:A1)..(xn:An)(I args)] with [I] an inductive definition;
returns [I] and [t'] or fails with a user error *)
-val reduce_to_quantified_ind : env -> evar_map -> types -> pinductive * types
+val reduce_to_quantified_ind : env -> evar_map -> types -> (inductive * EInstance.t) * types
(** [reduce_to_quantified_ref env sigma ref t] try to put [t] in the form
[t'=(x1:A1)..(xn:An)(ref args)] and fails with user error if not possible *)
@@ -91,7 +92,7 @@ val reduce_to_atomic_ref :
env -> evar_map -> global_reference -> types -> types
val find_hnf_rectype :
- env -> evar_map -> types -> pinductive * constr list
+ env -> evar_map -> types -> (inductive * EInstance.t) * constr list
val contextually : bool -> occurrences * constr_pattern ->
(patvar_map -> reduction_function) -> reduction_function
diff --git a/pretyping/typeclasses.ml b/pretyping/typeclasses.ml
index 01f3620f1..93c71e6ea 100644
--- a/pretyping/typeclasses.ml
+++ b/pretyping/typeclasses.ml
@@ -136,40 +136,40 @@ let typeclass_univ_instance (cl,u') =
let class_info c =
try Refmap.find c !classes
- with Not_found -> not_a_class (Global.env()) (printable_constr_of_global c)
+ with Not_found -> not_a_class (Global.env()) (EConstr.of_constr (printable_constr_of_global c))
-let global_class_of_constr env c =
- try let gr, u = Universes.global_of_constr c in
+let global_class_of_constr env sigma c =
+ try let gr, u = Termops.global_of_constr sigma c in
class_info gr, u
with Not_found -> not_a_class env c
-let dest_class_app env c =
- let cl, args = decompose_app c in
- global_class_of_constr env cl, args
+let dest_class_app env sigma c =
+ let cl, args = EConstr.decompose_app sigma c in
+ global_class_of_constr env sigma cl, (List.map EConstr.Unsafe.to_constr args)
-let dest_class_arity env c =
- let rels, c = decompose_prod_assum c in
- rels, dest_class_app env c
+let dest_class_arity env sigma c =
+ let open EConstr in
+ let rels, c = decompose_prod_assum sigma c in
+ rels, dest_class_app env sigma c
-let class_of_constr c =
- try Some (dest_class_arity (Global.env ()) c)
+let class_of_constr sigma c =
+ try Some (dest_class_arity (Global.env ()) sigma c)
with e when CErrors.noncritical e -> None
-let is_class_constr c =
- try let gr, u = Universes.global_of_constr c in
+let is_class_constr sigma c =
+ try let gr, u = Termops.global_of_constr sigma c in
Refmap.mem gr !classes
with Not_found -> false
let rec is_class_type evd c =
- let c, args = decompose_app c in
- match kind_of_term c with
+ let c, _ = Termops.decompose_app_vect evd c in
+ match EConstr.kind evd c with
| Prod (_, _, t) -> is_class_type evd t
- | Evar (e, _) when Evd.is_defined evd e ->
- is_class_type evd (Evarutil.whd_head_evar evd c)
- | _ -> is_class_constr c
+ | Cast (t, _, _) -> is_class_type evd t
+ | _ -> is_class_constr evd c
let is_class_evar evd evi =
- is_class_type evd evi.Evd.evar_concl
+ is_class_type evd (EConstr.of_constr evi.Evd.evar_concl)
(*
* classes persistent object
@@ -222,7 +222,7 @@ let discharge_class (_,cl) =
let discharge_context ctx' subst (grs, ctx) =
let grs' =
let newgrs = List.map (fun decl ->
- match decl |> RelDecl.get_type |> class_of_constr with
+ match decl |> RelDecl.get_type |> EConstr.of_constr |> class_of_constr Evd.empty with
| None -> None
| Some (_, ((tc,_), _)) -> Some (tc.cl_impl, true))
ctx'
@@ -284,15 +284,16 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
Nameops.add_suffix _id ("_subinstance_" ^ string_of_int !i))
in
let ty, ctx = Global.type_of_global_in_context env glob in
+ let ty = EConstr.of_constr ty in
let sigma = Evd.merge_context_set Evd.univ_rigid sigma (Univ.ContextSet.of_context ctx) in
let rec aux pri c ty path =
- let ty = Evarutil.nf_evar sigma ty in
- match class_of_constr ty with
+ match class_of_constr sigma ty with
| None -> []
| Some (rels, ((tc,u), args)) ->
let instapp =
- Reductionops.whd_beta sigma (appvectc c (Context.Rel.to_extended_vect 0 rels))
+ Reductionops.whd_beta sigma (EConstr.of_constr (appvectc c (Context.Rel.to_extended_vect mkRel 0 rels)))
in
+ let instapp = EConstr.Unsafe.to_constr instapp in
let projargs = Array.of_list (args @ [instapp]) in
let projs = List.map_filter
(fun (n, b, proj) ->
@@ -301,8 +302,10 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
| Some (Backward, _) -> None
| Some (Forward, info) ->
let proj = Option.get proj in
+ let rels = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) rels in
+ let u = EConstr.EInstance.kind sigma u in
let body = it_mkLambda_or_LetIn (mkApp (mkConstU (proj,u), projargs)) rels in
- if check && check_instance env sigma body then None
+ if check && check_instance env sigma (EConstr.of_constr body) then None
else
let newpri =
match pri, info.hint_priority with
@@ -314,7 +317,7 @@ let build_subclasses ~check env sigma glob { hint_priority = pri } =
in
let declare_proj hints (cref, info, body) =
let path' = cref :: path in
- let ty = Retyping.get_type_of env sigma body in
+ let ty = Retyping.get_type_of env sigma (EConstr.of_constr body) in
let rest = aux pri body ty path' in
hints @ (path', info, body) :: rest
in List.fold_left declare_proj [] projs
@@ -409,7 +412,7 @@ let remove_instance i =
let declare_instance info local glob =
let ty = Global.type_of_global_unsafe glob in
let info = Option.default {hint_priority = None; hint_pattern = None} info in
- match class_of_constr ty with
+ match class_of_constr Evd.empty (EConstr.of_constr ty) with
| Some (rels, ((tc,_), args) as _cl) ->
add_instance (new_instance tc info (not local) (Flags.use_polymorphic_flag ()) glob)
| None -> ()
diff --git a/pretyping/typeclasses.mli b/pretyping/typeclasses.mli
index 620bc367b..8d1c0b94c 100644
--- a/pretyping/typeclasses.mli
+++ b/pretyping/typeclasses.mli
@@ -61,13 +61,13 @@ val class_info : global_reference -> typeclass (** raises a UserError if not a c
(** These raise a UserError if not a class.
Caution: the typeclass structures is not instantiated w.r.t. the universe instance.
This is done separately by typeclass_univ_instance. *)
-val dest_class_app : env -> constr -> typeclass puniverses * constr list
+val dest_class_app : env -> evar_map -> EConstr.constr -> (typeclass * EConstr.EInstance.t) * constr list
(** Get the instantiated typeclass structure for a given universe instance. *)
val typeclass_univ_instance : typeclass puniverses -> typeclass puniverses
(** Just return None if not a class *)
-val class_of_constr : constr -> (Context.Rel.t * (typeclass puniverses * constr list)) option
+val class_of_constr : evar_map -> EConstr.constr -> (EConstr.rel_context * ((typeclass * EConstr.EInstance.t) * constr list)) option
val instance_impl : instance -> global_reference
@@ -99,11 +99,11 @@ val mark_unresolvables : ?filter:evar_filter -> evar_map -> evar_map
val mark_resolvables : ?filter:evar_filter -> evar_map -> evar_map
val mark_resolvable : evar_info -> evar_info
val is_class_evar : evar_map -> evar_info -> bool
-val is_class_type : evar_map -> types -> bool
+val is_class_type : evar_map -> EConstr.types -> bool
val resolve_typeclasses : ?fast_path:bool -> ?filter:evar_filter -> ?unique:bool ->
?split:bool -> ?fail:bool -> env -> evar_map -> evar_map
-val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> types -> open_constr
+val resolve_one_typeclass : ?unique:bool -> env -> evar_map -> EConstr.types -> evar_map * EConstr.constr
val set_typeclass_transparency_hook : (evaluable_global_reference -> bool (*local?*) -> bool -> unit) Hook.t
val set_typeclass_transparency : evaluable_global_reference -> bool -> bool -> unit
@@ -120,7 +120,7 @@ val add_instance_hint : global_reference_or_constr -> global_reference list ->
val remove_instance_hint : global_reference -> unit
val solve_all_instances_hook : (env -> evar_map -> evar_filter -> bool -> bool -> bool -> evar_map) Hook.t
-val solve_one_instance_hook : (env -> evar_map -> types -> bool -> open_constr) Hook.t
+val solve_one_instance_hook : (env -> evar_map -> EConstr.types -> bool -> evar_map * EConstr.constr) Hook.t
val declare_instance : Vernacexpr.hint_info_expr option -> bool -> global_reference -> unit
diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml
index b1dfb19a0..2db0e9e88 100644
--- a/pretyping/typeclasses_errors.ml
+++ b/pretyping/typeclasses_errors.ml
@@ -9,6 +9,7 @@
(*i*)
open Names
open Term
+open EConstr
open Environ
open Constrexpr
open Globnames
diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli
index ee76f6383..9bd430e4d 100644
--- a/pretyping/typeclasses_errors.mli
+++ b/pretyping/typeclasses_errors.mli
@@ -9,6 +9,7 @@
open Loc
open Names
open Term
+open EConstr
open Environ
open Constrexpr
open Globnames
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index e79e3d46f..c2a030bcd 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -6,47 +6,58 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Term
-open Vars
open Environ
+open EConstr
+open Vars
open Reductionops
-open Type_errors
open Inductive
open Inductiveops
open Typeops
open Arguments_renaming
+open Pretype_errors
open Context.Rel.Declaration
+let push_rec_types pfix env =
+ let (i, c, t) = pfix in
+ let inj c = EConstr.Unsafe.to_constr c in
+ push_rec_types (i, Array.map inj c, Array.map inj t) env
+
let meta_type evd mv =
let ty =
try Evd.meta_ftype evd mv
with Not_found -> anomaly (str "unknown meta ?" ++ str (Nameops.string_of_meta mv)) in
+ let ty = Evd.map_fl EConstr.of_constr ty in
meta_instance evd ty
-let constant_type_knowing_parameters env cst jl =
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
- type_of_constant_knowing_parameters_in env cst paramstyp
+let constant_type_knowing_parameters env sigma (cst, u) jl =
+ let u = Unsafe.to_instance u in
+ let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in
+ EConstr.of_constr (type_of_constant_knowing_parameters_in env (cst, u) paramstyp)
-let inductive_type_knowing_parameters env (ind,u) jl =
+let inductive_type_knowing_parameters env sigma (ind,u) jl =
+ let u = Unsafe.to_instance u in
let mspec = lookup_mind_specif env ind in
- let paramstyp = Array.map (fun j -> lazy j.uj_type) jl in
- Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp
+ let paramstyp = Array.map (fun j -> lazy (EConstr.to_constr sigma j.uj_type)) jl in
+ EConstr.of_constr (Inductive.type_of_inductive_knowing_parameters env (mspec,u) paramstyp)
let e_type_judgment env evdref j =
- match kind_of_term (whd_all env !evdref j.uj_type) with
- | Sort s -> {utj_val = j.uj_val; utj_type = s }
+ match EConstr.kind !evdref (whd_all env !evdref j.uj_type) with
+ | Sort s -> {utj_val = j.uj_val; utj_type = ESorts.kind !evdref s }
| Evar ev ->
let (evd,s) = Evardefine.define_evar_as_sort env !evdref ev in
evdref := evd; { utj_val = j.uj_val; utj_type = s }
- | _ -> error_not_type env j
+ | _ -> error_not_a_type env !evdref j
let e_assumption_of_judgment env evdref j =
try (e_type_judgment env evdref j).utj_val
- with TypeError _ ->
- error_assumption env j
+ with Type_errors.TypeError _ | PretypeError _ ->
+ error_assumption env !evdref j
let e_judge_of_apply env evdref funj argjv =
let rec apply_rec n typ = function
@@ -54,28 +65,28 @@ let e_judge_of_apply env evdref funj argjv =
{ uj_val = mkApp (j_val funj, Array.map j_val argjv);
uj_type = typ }
| hj::restjl ->
- match kind_of_term (whd_all env !evdref typ) with
+ match EConstr.kind !evdref (whd_all env !evdref typ) with
| Prod (_,c1,c2) ->
if Evarconv.e_cumul env evdref hj.uj_type c1 then
apply_rec (n+1) (subst1 hj.uj_val c2) restjl
else
- error_cant_apply_bad_type env (n,c1, hj.uj_type) funj argjv
+ error_cant_apply_bad_type env !evdref (n, c1, hj.uj_type) funj argjv
| Evar ev ->
let (evd',t) = Evardefine.define_evar_as_product !evdref ev in
evdref := evd';
- let (_,_,c2) = destProd t in
+ let (_,_,c2) = destProd evd' t in
apply_rec (n+1) (subst1 hj.uj_val c2) restjl
| _ ->
- error_cant_apply_not_functional env funj argjv
+ error_cant_apply_not_functional env !evdref funj argjv
in
apply_rec 1 funj.uj_type (Array.to_list argjv)
let e_check_branch_types env evdref (ind,u) cj (lfj,explft) =
if not (Int.equal (Array.length lfj) (Array.length explft)) then
- error_number_branches env cj (Array.length explft);
+ error_number_branches env !evdref cj (Array.length explft);
for i = 0 to Array.length explft - 1 do
if not (Evarconv.e_cumul env evdref lfj.(i).uj_type explft.(i)) then
- error_ill_formed_branch env cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
+ error_ill_formed_branch env !evdref cj.uj_val ((ind,i+1),u) lfj.(i).uj_type explft.(i)
done
let max_sort l =
@@ -83,21 +94,22 @@ let max_sort l =
if Sorts.List.mem InSet l then InSet else InProp
let e_is_correct_arity env evdref c pj ind specif params =
- let arsign = make_arity_signature env true (make_ind_family (ind,params)) in
+ let arsign = make_arity_signature env !evdref true (make_ind_family (ind,params)) in
let allowed_sorts = elim_sorts specif in
- let error () = error_elim_arity env ind allowed_sorts c pj None in
+ let error () = Pretype_errors.error_elim_arity env !evdref ind allowed_sorts c pj None in
let rec srec env pt ar =
let pt' = whd_all env !evdref pt in
- match kind_of_term pt', ar with
+ match EConstr.kind !evdref pt', ar with
| Prod (na1,a1,t), (LocalAssum (_,a1'))::ar' ->
if not (Evarconv.e_cumul env evdref a1 a1') then error ();
srec (push_rel (LocalAssum (na1,a1)) env) t ar'
| Sort s, [] ->
+ let s = ESorts.kind !evdref s in
if not (Sorts.List.mem (Sorts.family s) allowed_sorts)
then error ()
| Evar (ev,_), [] ->
let evd, s = Evd.fresh_sort_in_family env !evdref (max_sort allowed_sorts) in
- evdref := Evd.define ev (mkSort s) evd
+ evdref := Evd.define ev (Constr.mkSort s) evd
| _, (LocalDef _ as d)::ar' ->
srec (push_rel d env) (lift 1 pt') ar'
| _ ->
@@ -105,21 +117,35 @@ let e_is_correct_arity env evdref c pj ind specif params =
in
srec env pj.uj_type (List.rev arsign)
+let lambda_applist_assum sigma n c l =
+ let rec app n subst t l =
+ if Int.equal n 0 then
+ if l == [] then substl subst t
+ else anomaly (Pp.str "Not enough arguments")
+ else match EConstr.kind sigma t, l with
+ | Lambda(_,_,c), arg::l -> app (n-1) (arg::subst) c l
+ | LetIn(_,b,_,c), _ -> app (n-1) (substl subst b::subst) c l
+ | _ -> anomaly (Pp.str "Not enough lambda/let's") in
+ app n [] c l
+
let e_type_case_branches env evdref (ind,largs) pj c =
let specif = lookup_mind_specif env (fst ind) in
let nparams = inductive_params specif in
let (params,realargs) = List.chop nparams largs in
let p = pj.uj_val in
+ let params = List.map EConstr.Unsafe.to_constr params in
let () = e_is_correct_arity env evdref c pj ind specif params in
- let lc = build_branches_type ind specif params p in
+ let lc = build_branches_type ind specif params (EConstr.to_constr !evdref p) in
+ let lc = Array.map EConstr.of_constr lc in
let n = (snd specif).Declarations.mind_nrealdecls in
- let ty = whd_betaiota !evdref (lambda_applist_assum (n+1) p (realargs@[c])) in
+ let ty = whd_betaiota !evdref (lambda_applist_assum !evdref (n+1) p (realargs@[c])) in
(lc, ty)
let e_judge_of_case env evdref ci pj cj lfj =
- let indspec =
+ let ((ind, u), spec) =
try find_mrectype env !evdref cj.uj_type
- with Not_found -> error_case_not_inductive env cj in
+ with Not_found -> error_case_not_inductive env !evdref cj in
+ let indspec = ((ind, EInstance.kind !evdref u), spec) in
let _ = check_case_info env (fst indspec) ci in
let (bty,rslty) = e_type_case_branches env evdref indspec pj cj.uj_val in
e_check_branch_types env evdref (fst indspec) cj (lfj,bty);
@@ -132,25 +158,25 @@ let check_type_fixpoint loc env evdref lna lar vdefj =
for i = 0 to lt-1 do
if not (Evarconv.e_cumul env evdref (vdefj.(i)).uj_type
(lift lt lar.(i))) then
- Pretype_errors.error_ill_typed_rec_body ~loc env !evdref
+ error_ill_typed_rec_body ~loc env !evdref
i lna vdefj lar
done
(* 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 (sort_of_arity env sigma pj.uj_type) in
+ let ksort = family_of_sort (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
let s = inductive_sort_family (snd specif) in
- error_elim_arity env ind sorts c pj
- (Some(ksort,s,error_elim_explain ksort s))
+ error_elim_arity env sigma ind sorts c pj
+ (Some(ksort,s,Type_errors.error_elim_explain ksort s))
let e_judge_of_cast env evdref cj k tj =
let expected_type = tj.utj_val in
if not (Evarconv.e_cumul env evdref cj.uj_type expected_type) then
- error_actual_type env cj expected_type;
+ error_actual_type_core env !evdref cj expected_type;
{ uj_val = mkCast (cj.uj_val, k, expected_type);
uj_type = expected_type }
@@ -159,19 +185,78 @@ let enrich_env env evdref =
let penv' = Pre_env.({ penv with env_stratification =
{ penv.env_stratification with env_universes = Evd.universes !evdref } }) in
Environ.env_of_pre_env penv'
-
+
+let check_fix env sigma pfix =
+ let inj c = EConstr.to_constr sigma c in
+ let (idx, (ids, cs, ts)) = pfix in
+ check_fix env (idx, (ids, Array.map inj cs, Array.map inj ts))
+
+let check_cofix env sigma pcofix =
+ let inj c = EConstr.to_constr sigma c in
+ let (idx, (ids, cs, ts)) = pcofix in
+ check_cofix env (idx, (ids, Array.map inj cs, Array.map inj ts))
+
(* The typing machine with universes and existential variables. *)
+let judge_of_prop =
+ { uj_val = EConstr.mkProp;
+ uj_type = EConstr.mkSort type1_sort }
+
+let judge_of_set =
+ { uj_val = EConstr.mkSet;
+ uj_type = EConstr.mkSort type1_sort }
+
+let judge_of_prop_contents = function
+ | Null -> judge_of_prop
+ | Pos -> judge_of_set
+
+let judge_of_type u =
+ let uu = Univ.Universe.super u in
+ { uj_val = EConstr.mkType u;
+ uj_type = EConstr.mkType uu }
+
+let judge_of_relative env v =
+ Termops.on_judgment EConstr.of_constr (judge_of_relative env v)
+
+let judge_of_variable env id =
+ Termops.on_judgment EConstr.of_constr (judge_of_variable env id)
+
+let judge_of_projection env sigma p cj =
+ let pb = lookup_projection p env in
+ let (ind,u), args =
+ try find_mrectype env sigma cj.uj_type
+ with Not_found -> error_case_not_inductive env sigma cj
+ in
+ let u = EInstance.kind sigma u in
+ let ty = EConstr.of_constr (CVars.subst_instance_constr u pb.Declarations.proj_type) in
+ let ty = substl (cj.uj_val :: List.rev args) ty in
+ {uj_val = EConstr.mkProj (p,cj.uj_val);
+ uj_type = ty}
+
+let judge_of_abstraction env name var j =
+ { uj_val = mkLambda (name, var.utj_val, j.uj_val);
+ uj_type = mkProd (name, var.utj_val, j.uj_type) }
+
+let judge_of_product env name t1 t2 =
+ let s = sort_of_product env t1.utj_type t2.utj_type in
+ { uj_val = mkProd (name, t1.utj_val, t2.utj_val);
+ uj_type = mkSort s }
+
+let judge_of_letin env name defj typj j =
+ { uj_val = mkLetIn (name, defj.uj_val, typj.utj_val, j.uj_val) ;
+ uj_type = subst1 defj.uj_val j.uj_type }
+
(* cstr must be in n.f. w.r.t. evars and execute returns a judgement
where both the term and type are in n.f. *)
let rec execute env evdref cstr =
- match kind_of_term cstr with
+ let cstr = whd_evar !evdref cstr in
+ match EConstr.kind !evdref cstr with
| Meta n ->
{ uj_val = cstr; uj_type = meta_type !evdref n }
| Evar ev ->
- let ty = Evd.existential_type !evdref ev in
- let jty = execute env evdref (whd_evar !evdref ty) in
+ let ty = EConstr.existential_type !evdref ev in
+ let jty = execute env evdref ty in
let jty = e_assumption_of_judgment env evdref jty in
{ uj_val = cstr; uj_type = jty }
@@ -181,14 +266,17 @@ let rec execute env evdref cstr =
| Var id ->
judge_of_variable env id
- | Const c ->
- make_judge cstr (rename_type_of_constant env c)
+ | Const (c, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_constant env (c, u)))
- | Ind ind ->
- make_judge cstr (rename_type_of_inductive env ind)
+ | Ind (ind, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_inductive env (ind, u)))
- | Construct cstruct ->
- make_judge cstr (rename_type_of_constructor env cstruct)
+ | Construct (cstruct, u) ->
+ let u = EInstance.kind !evdref u in
+ make_judge cstr (EConstr.of_constr (rename_type_of_constructor env (cstruct, u)))
| Case (ci,p,c,lf) ->
let cj = execute env evdref c in
@@ -199,39 +287,39 @@ let rec execute env evdref cstr =
| Fix ((vn,i as vni),recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let fix = (vni,recdef') in
- check_fix env fix;
+ check_fix env !evdref fix;
make_judge (mkFix fix) tys.(i)
| CoFix (i,recdef) ->
let (_,tys,_ as recdef') = execute_recdef env evdref recdef in
let cofix = (i,recdef') in
- check_cofix env cofix;
+ check_cofix env !evdref cofix;
make_judge (mkCoFix cofix) tys.(i)
- | Sort (Prop c) ->
- judge_of_prop_contents c
-
- | Sort (Type u) ->
+ | Sort s ->
+ begin match ESorts.kind !evdref s with
+ | Prop c ->
+ judge_of_prop_contents c
+ | Type u ->
judge_of_type u
+ end
| Proj (p, c) ->
let cj = execute env evdref c in
- judge_of_projection env p (Evarutil.j_nf_evar !evdref cj)
+ judge_of_projection env !evdref p cj
| App (f,args) ->
let jl = execute_array env evdref args in
let j =
- match kind_of_term f with
- | Ind ind when Environ.template_polymorphic_pind ind env ->
+ match EConstr.kind !evdref f with
+ | Ind (ind, u) when EInstance.is_empty u && Environ.template_polymorphic_ind ind env ->
(* Sort-polymorphism of inductive types *)
make_judge f
- (inductive_type_knowing_parameters env ind
- (Evarutil.jv_nf_evar !evdref jl))
- | Const cst when Environ.template_polymorphic_pconstant cst env ->
+ (inductive_type_knowing_parameters env !evdref (ind, u) jl)
+ | Const (cst, u) when EInstance.is_empty u && Environ.template_polymorphic_constant cst env ->
(* Sort-polymorphism of inductive types *)
make_judge f
- (constant_type_knowing_parameters env cst
- (Evarutil.jv_nf_evar !evdref jl))
+ (constant_type_knowing_parameters env !evdref (cst, u) jl)
| _ ->
execute env evdref f
in
@@ -282,7 +370,7 @@ let e_check env evdref c t =
let env = enrich_env env evdref in
let j = execute env evdref c in
if not (Evarconv.e_cumul env evdref j.uj_type t) then
- error_actual_type env j (nf_evar !evdref t)
+ error_actual_type_core env !evdref j t
(* Type of a constr *)
@@ -327,4 +415,4 @@ let e_solve_evars env evdref c =
(* side-effect on evdref *)
nf_evar !evdref c
-let _ = Evarconv.set_solve_evars e_solve_evars
+let _ = Evarconv.set_solve_evars (fun env evdref c -> e_solve_evars env evdref c)
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 04e5e40bc..91134b499 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -6,8 +6,10 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+open Names
open Term
open Environ
+open EConstr
open Evd
(** This module provides the typing machine with existential variables
@@ -44,3 +46,10 @@ val check_allowed_sort : env -> evar_map -> pinductive -> constr -> constr ->
expected ones *)
val check_type_fixpoint : Loc.t -> env -> evar_map ref ->
Names.Name.t array -> types array -> unsafe_judgment array -> unit
+
+val judge_of_prop : unsafe_judgment
+val judge_of_set : unsafe_judgment
+val judge_of_abstraction : Environ.env -> Name.t ->
+ unsafe_type_judgment -> unsafe_judgment -> unsafe_judgment
+val judge_of_product : Environ.env -> Name.t ->
+ unsafe_type_judgment -> unsafe_type_judgment -> unsafe_judgment
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
index 11771f7ba..532cc8baa 100644
--- a/pretyping/unification.ml
+++ b/pretyping/unification.ml
@@ -6,15 +6,18 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open CErrors
open Pp
open Util
open Names
open Term
-open Vars
open Termops
-open Namegen
open Environ
+open EConstr
+open Vars
+open Namegen
open Evd
open Reduction
open Reductionops
@@ -30,6 +33,13 @@ open Locusops
open Find_subterm
open Sigma.Notations
+type metabinding = (metavariable * EConstr.constr * (instance_constraint * instance_typing_status))
+
+type subst0 =
+ (evar_map *
+ metabinding list *
+ (Environ.env * EConstr.existential * EConstr.t) list)
+
module RelDecl = Context.Rel.Declaration
module NamedDecl = Context.Named.Declaration
@@ -54,7 +64,22 @@ let _ = Goptions.declare_bool_option {
Goptions.optwrite = (fun a -> debug_unification:=a);
}
+(** Making this unification algorithm correct w.r.t. the evar-map abstraction
+ breaks too much stuff. So we redefine incorrect functions here. *)
+
+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
+ | Evar _ -> raise Occur
+ | Meta _ -> raise Occur
+ | _ -> iter_constr occrec c
+ in try occrec c; false with Occur -> true
+
+
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
| Meta _ -> raise Occur
| Evar (ev,args) ->
@@ -68,26 +93,29 @@ let occur_meta_or_undefined_evar evd c =
let occur_meta_evd sigma mv c =
let rec occrec c =
(* Note: evars are not instantiated by terms with metas *)
- let c = whd_evar sigma (whd_meta sigma c) in
- match kind_of_term c with
+ let c = whd_meta sigma c in
+ match EConstr.kind sigma c with
| Meta mv' when Int.equal mv mv' -> raise Occur
- | _ -> Constr.iter occrec c
+ | _ -> EConstr.iter sigma occrec c
in try occrec c; false with Occur -> true
(* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms,
gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *)
let abstract_scheme env evd c l lname_typ =
+ let mkLambda_name env (n,a,b) =
+ mkLambda (named_hd env evd a n, a, b)
+ in
List.fold_left2
(fun (t,evd) (locc,a) decl ->
let na = RelDecl.get_name decl in
let ta = RelDecl.get_type decl in
- let na = match kind_of_term a with Var id -> Name id | _ -> na in
+ let na = match EConstr.kind evd a with Var id -> Name id | _ -> na in
(* [occur_meta ta] test removed for support of eelim/ecase but consequences
are unclear...
if occur_meta ta then error "cannot find a type for the generalisation"
else *)
- if occur_meta a then mkLambda_name env (na,ta,t), evd
+ if occur_meta evd a then mkLambda_name env (na,ta,t), evd
else
let t', evd' = Find_subterm.subst_closed_term_occ env evd locc a t in
mkLambda_name env (na,ta,t'), evd')
@@ -107,6 +135,9 @@ let abstract_list_all env evd typ c l =
| UserError _ ->
error_cannot_find_well_typed_abstraction env evd p l None
| Type_errors.TypeError (env',x) ->
+ (** FIXME: plug back the typing information *)
+ error_cannot_find_well_typed_abstraction env evd p l None
+ | Pretype_errors.PretypeError (env',evd,TypingError x) ->
error_cannot_find_well_typed_abstraction env evd p l (Some (env',x)) in
evd,(p,typp)
@@ -117,17 +148,17 @@ let abstract_list_all_with_dependencies env evd typ c l =
let evd = Sigma.Unsafe.of_evar_map evd in
let Sigma (ev, evd, _) = new_evar env evd typ in
let evd = Sigma.to_evar_map evd in
- let evd,ev' = evar_absorb_arguments env evd (destEvar ev) l in
+ let evd,ev' = evar_absorb_arguments env evd (destEvar evd ev) l in
let n = List.length l in
let argoccs = set_occurrences_of_last_arg (Array.sub (snd ev') 0 n) in
let evd,b =
Evarconv.second_order_matching empty_transparent_state
env evd ev' argoccs c in
if b then
- let p = nf_evar evd (existential_value evd (destEvar ev)) in
+ let p = nf_evar evd ev in
evd, p
else error_cannot_find_well_typed_abstraction env evd
- (nf_evar evd c) l None
+ c l None
(**)
@@ -147,50 +178,53 @@ let extract_instance_status = function
| CUMUL -> add_type_status (IsSubType, IsSuperType)
| CONV -> add_type_status (Conv, Conv)
-let rec subst_meta_instances bl c =
- match kind_of_term c with
+let rec subst_meta_instances sigma bl c =
+ match EConstr.kind sigma c with
| Meta i ->
let select (j,_,_) = Int.equal i j in
(try pi2 (List.find select bl) with Not_found -> c)
- | _ -> Constr.map (subst_meta_instances bl) c
+ | _ -> EConstr.map sigma (subst_meta_instances sigma bl) c
(** [env] should be the context in which the metas live *)
let pose_all_metas_as_evars env evd t =
let evdref = ref evd in
- let rec aux t = match kind_of_term t with
+ let rec aux t = match EConstr.kind !evdref t with
| Meta mv ->
(match Evd.meta_opt_fvalue !evdref mv with
- | Some ({rebus=c},_) -> c
+ | Some ({rebus=c},_) -> EConstr.of_constr c
| None ->
let {rebus=ty;freemetas=mvs} = Evd.meta_ftype evd mv in
+ let ty = EConstr.of_constr ty in
let ty = if Evd.Metaset.is_empty mvs then ty else aux ty in
let src = Evd.evar_source_of_meta mv !evdref in
let ev = Evarutil.e_new_evar env evdref ~src ty in
- evdref := meta_assign mv (ev,(Conv,TypeNotProcessed)) !evdref;
+ evdref := meta_assign mv (EConstr.Unsafe.to_constr ev,(Conv,TypeNotProcessed)) !evdref;
ev)
| _ ->
- Constr.map aux t in
+ EConstr.map !evdref aux t in
let c = aux t in
(* side-effect *)
(!evdref, c)
-let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst) =
- match kind_of_term f with
+let solve_pattern_eqn_array (env,nb) f l c (sigma,metasubst,evarsubst : subst0) =
+ match EConstr.kind sigma f with
| Meta k ->
(* We enforce that the Meta does not depend on the [nb]
extra assumptions added by unification to the context *)
let env' = pop_rel_context nb env in
let sigma,c = pose_all_metas_as_evars env' sigma c in
- let c = solve_pattern_eqn env l c in
+ let c = solve_pattern_eqn env sigma l c in
let pb = (Conv,TypeNotProcessed) in
- if noccur_between 1 nb c then
+ if noccur_between sigma 1 nb c then
sigma,(k,lift (-nb) c,pb)::metasubst,evarsubst
- else error_cannot_unify_local env sigma (applist (f, l),c,c)
+ else
+ let l = List.map of_alias l in
+ error_cannot_unify_local env sigma (applist (f, l),c,c)
| Evar ev ->
let env' = pop_rel_context nb env in
let sigma,c = pose_all_metas_as_evars env' sigma c in
- sigma,metasubst,(env,ev,solve_pattern_eqn env l c)::evarsubst
+ sigma,metasubst,(env,ev,solve_pattern_eqn env sigma l c)::evarsubst
| _ -> assert false
let push d (env,n) = (push_rel_assum d env,n+1)
@@ -454,15 +488,15 @@ let use_evars_pattern_unification flags =
!global_pattern_unification_flag && flags.use_pattern_unification
&& Flags.version_strictly_greater Flags.V8_2
-let use_metas_pattern_unification flags nb l =
+let use_metas_pattern_unification sigma flags nb l =
!global_pattern_unification_flag && flags.use_pattern_unification
|| (Flags.version_less_or_equal Flags.V8_3 ||
flags.use_meta_bound_pattern_unification) &&
- Array.for_all (fun c -> isRel c && destRel c <= nb) l
+ Array.for_all (fun c -> isRel sigma c && destRel sigma c <= nb) l
type key =
| IsKey of CClosure.table_key
- | IsProj of projection * constr
+ | IsProj of projection * EConstr.constr
let expand_table_key env = function
| ConstKey cst -> constant_opt_value_in env cst
@@ -478,11 +512,11 @@ let unfold_projection env p stk =
| None -> assert false)
let expand_key ts env sigma = function
- | Some (IsKey k) -> expand_table_key env k
+ | Some (IsKey k) -> Option.map EConstr.of_constr (expand_table_key env k)
| Some (IsProj (p, c)) ->
- let red = Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma
+ let red = Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state ts env sigma
Cst_stack.empty (c, unfold_projection env p [])))
- in if Term.eq_constr (mkProj (p, c)) red then None else Some red
+ in if EConstr.eq_constr sigma (EConstr.mkProj (p, c)) red then None else Some red
| None -> None
@@ -495,12 +529,13 @@ type unirec_flags = {
let subterm_restriction opt flags =
not opt.at_top && flags.restrict_conv_on_strict_subterms
-let key_of env b flags f =
+let key_of env sigma b flags f =
if subterm_restriction b flags then None else
- match kind_of_term f with
+ match EConstr.kind sigma f with
| Const (cst, u) when is_transparent env (ConstKey cst) &&
(Cpred.mem cst (snd flags.modulo_delta)
|| Environ.is_projection cst env) ->
+ let u = EInstance.kind sigma u in
Some (IsKey (ConstKey (cst, u)))
| Var id when is_transparent env (VarKey id) &&
Id.Pred.mem id (fst flags.modulo_delta) ->
@@ -542,8 +577,8 @@ let oracle_order env cf1 cf2 =
Some (Conv_oracle.oracle_order (fun x -> x)
(Environ.oracle env) false (translate_key k1) (translate_key k2))
-let is_rigid_head flags t =
- match kind_of_term t with
+let is_rigid_head sigma flags t =
+ match EConstr.kind sigma t with
| Const (cst,u) -> not (Cpred.mem cst (snd flags.modulo_delta))
| Ind (i,u) -> true
| Construct _ -> true
@@ -559,15 +594,15 @@ let force_eqs c =
let constr_cmp pb sigma flags t u =
let cstrs =
- if pb == Reduction.CONV then Universes.eq_constr_universes t u
- else Universes.leq_constr_universes t u
+ if pb == Reduction.CONV then EConstr.eq_constr_universes sigma t u
+ else EConstr.leq_constr_universes sigma t u
in
match cstrs with
| Some cstrs ->
begin try Evd.add_universe_constraints sigma cstrs, true
with Univ.UniverseInconsistency _ -> sigma, false
| Evd.UniversesDiffer ->
- if is_rigid_head flags t then
+ if is_rigid_head sigma flags t then
try Evd.add_universe_constraints sigma (force_eqs cstrs), true
with Univ.UniverseInconsistency _ -> sigma, false
else sigma, false
@@ -576,34 +611,38 @@ let constr_cmp pb sigma flags t u =
sigma, false
let do_reduce ts (env, nb) sigma c =
- Stack.zip (fst (whd_betaiota_deltazeta_for_iota_state
+ Stack.zip sigma (fst (whd_betaiota_deltazeta_for_iota_state
ts env sigma Cst_stack.empty (c, Stack.empty)))
let use_full_betaiota flags =
flags.modulo_betaiota && Flags.version_strictly_greater Flags.V8_3
-let isAllowedEvar flags c = match kind_of_term c with
+let isAllowedEvar sigma flags c = match EConstr.kind sigma c with
| Evar (evk,_) -> not (Evar.Set.mem evk flags.frozen_evars)
| _ -> false
-let subst_defined_metas_evars (bl,el) c =
+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
| Meta i ->
let select (j,_,_) = Int.equal i j in
- substrec (pi2 (List.find select bl))
+ substrec (EConstr.Unsafe.to_constr (pi2 (List.find select bl)))
| Evar (evk,args) ->
- let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.equal Constr.equal args args' in
- (try substrec (pi3 (List.find select el))
+ let eq c1 c2 = Constr.equal c1 (EConstr.Unsafe.to_constr c2) in
+ let select (_,(evk',args'),_) = Evar.equal evk evk' && Array.for_all2 eq args args' in
+ (try substrec (EConstr.Unsafe.to_constr (pi3 (List.find select el)))
with Not_found -> Constr.map substrec c)
| _ -> Constr.map substrec c
- in try Some (substrec c) with Not_found -> None
+ in try Some (EConstr.of_constr (substrec c)) with Not_found -> None
-let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN =
- match subst_defined_metas_evars (metasubst,[]) tyM with
+let check_compatibility env pbty flags (sigma,metasubst,evarsubst : subst0) tyM tyN =
+ match subst_defined_metas_evars sigma (metasubst,[]) tyM with
| None -> sigma
| Some m ->
- match subst_defined_metas_evars (metasubst,[]) tyN with
+ match subst_defined_metas_evars sigma (metasubst,[]) tyN with
| None -> sigma
| Some n ->
if is_ground_term sigma m && is_ground_term sigma n then
@@ -613,9 +652,9 @@ let check_compatibility env pbty flags (sigma,metasubst,evarsubst) tyM tyN =
else sigma
-let rec is_neutral env ts t =
- let (f, l) = decompose_appvect t in
- match kind_of_term f with
+let rec is_neutral env sigma ts t =
+ let (f, l) = decompose_app_vect sigma t in
+ match EConstr.kind sigma f with
| Const (c, u) ->
not (Environ.evaluable_constant c env) ||
not (is_transparent env (ConstKey c)) ||
@@ -626,24 +665,24 @@ let rec is_neutral env ts t =
not (Id.Pred.mem id (fst ts))
| Rel n -> true
| Evar _ | Meta _ -> true
- | Case (_, p, c, cl) -> is_neutral env ts c
- | Proj (p, c) -> is_neutral env ts c
+ | Case (_, p, c, cl) -> is_neutral env sigma ts c
+ | Proj (p, c) -> is_neutral env sigma ts c
| _ -> false
-let is_eta_constructor_app env ts f l1 term =
- match kind_of_term f with
+let is_eta_constructor_app env sigma ts f l1 term =
+ match EConstr.kind sigma f with
| Construct (((_, i as ind), j), u) when i == 0 && j == 1 ->
let mib = lookup_mind (fst ind) env in
(match mib.Declarations.mind_record with
| Some (Some (_,exp,projs)) when mib.Declarations.mind_finite == Decl_kinds.BiFinite &&
Array.length projs == Array.length l1 - mib.Declarations.mind_nparams ->
(** Check that the other term is neutral *)
- is_neutral env ts term
+ is_neutral env sigma ts term
| _ -> false)
| _ -> false
-let eta_constructor_app env f l1 term =
- match kind_of_term f with
+let eta_constructor_app env sigma f l1 term =
+ match EConstr.kind sigma f with
| Construct (((_, i as ind), j), u) ->
let mib = lookup_mind (fst ind) env in
(match mib.Declarations.mind_record with
@@ -656,15 +695,15 @@ let eta_constructor_app env f l1 term =
| _ -> assert false)
| _ -> assert false
-let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb flags m n =
- let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn) curm curn =
+let rec unify_0_with_initial_metas (sigma,ms,es as subst : subst0) conv_at_top env cv_pb flags m n =
+ let rec unirec_rec (curenv,nb as curenvnb) pb opt ((sigma,metasubst,evarsubst) as substn : subst0) curm curn =
let cM = Evarutil.whd_head_evar sigma curm
and cN = Evarutil.whd_head_evar sigma curn in
let () =
if !debug_unification then
- Feedback.msg_debug (Termops.print_constr_env curenv cM ++ str" ~= " ++ Termops.print_constr_env curenv cN)
+ Feedback.msg_debug (print_constr_env curenv sigma cM ++ str" ~= " ++ print_constr_env curenv sigma cN)
in
- match (kind_of_term cM,kind_of_term cN) with
+ match (EConstr.kind sigma cM, EConstr.kind sigma cN) with
| Meta k1, Meta k2 ->
if Int.equal k1 k2 then substn else
let stM,stN = extract_instance_status pb in
@@ -679,7 +718,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
if k2 < k1 then sigma,(k1,cN,stN)::metasubst,evarsubst
else sigma,(k2,cM,stM)::metasubst,evarsubst
| Meta k, _
- when not (dependent cM cN) (* helps early trying alternatives *) ->
+ when not (dependent sigma cM cN) (* helps early trying alternatives *) ->
let sigma =
if opt.with_types && flags.check_applied_meta_types then
(try
@@ -693,13 +732,13 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(* Here we check that [cN] does not contain any local variables *)
if Int.equal nb 0 then
sigma,(k,cN,snd (extract_instance_status pb))::metasubst,evarsubst
- else if noccur_between 1 nb cN then
+ else if noccur_between sigma 1 nb cN then
(sigma,
(k,lift (-nb) cN,snd (extract_instance_status pb))::metasubst,
evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| _, Meta k
- when not (dependent cN cM) (* helps early trying alternatives *) ->
+ when not (dependent sigma cN cM) (* helps early trying alternatives *) ->
let sigma =
if opt.with_types && flags.check_applied_meta_types then
(try
@@ -713,7 +752,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(* Here we check that [cM] does not contain any local variables *)
if Int.equal nb 0 then
(sigma,(k,cM,fst (extract_instance_status pb))::metasubst,evarsubst)
- else if noccur_between 1 nb cM
+ else if noccur_between sigma 1 nb cM
then
(sigma,(k,lift (-nb) cM,fst (extract_instance_status pb))::metasubst,
evarsubst)
@@ -728,20 +767,22 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
sigma,metasubst,((curenv,ev,cN)::evarsubst)
| Evar (evk,_ as ev), _
when not (Evar.Set.mem evk flags.frozen_evars)
- && not (occur_evar evk cN) ->
- let cmvars = free_rels cM and cnvars = free_rels cN in
+ && not (occur_evar sigma evk cN) ->
+ let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in
if Int.Set.subset cnvars cmvars then
sigma,metasubst,((curenv,ev,cN)::evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| _, Evar (evk,_ as ev)
when not (Evar.Set.mem evk flags.frozen_evars)
- && not (occur_evar evk cM) ->
- let cmvars = free_rels cM and cnvars = free_rels cN in
+ && not (occur_evar sigma evk cM) ->
+ let cmvars = free_rels sigma cM and cnvars = free_rels sigma cN in
if Int.Set.subset cmvars cnvars then
sigma,metasubst,((curenv,ev,cM)::evarsubst)
else error_cannot_unify_local curenv sigma (m,n,cN)
| Sort s1, Sort s2 ->
(try
+ let s1 = ESorts.kind sigma s1 in
+ let s2 = ESorts.kind sigma s2 in
let sigma' =
if pb == CUMUL
then Evd.set_leq_sort curenv sigma s1 s2
@@ -779,30 +820,30 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| App (f1, l1), _ when flags.modulo_eta &&
(* This ensures cN is an evar, meta or irreducible constant/variable
and not a constructor. *)
- is_eta_constructor_app curenv flags.modulo_delta f1 l1 cN ->
+ is_eta_constructor_app curenv sigma flags.modulo_delta f1 l1 cN ->
(try
- let l1', l2' = eta_constructor_app curenv f1 l1 cN in
+ let l1', l2' = eta_constructor_app curenv sigma f1 l1 cN in
let opt' = {opt with at_top = true; with_cs = false} in
Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2'
with ex when precatchable_exception ex ->
- match kind_of_term cN with
+ match EConstr.kind sigma cN with
| App(f2,l2) when
- (isMeta f2 && use_metas_pattern_unification flags nb l2
- || use_evars_pattern_unification flags && isAllowedEvar flags f2) ->
+ (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) ->
unify_app_pattern false curenvnb pb opt substn cM f1 l1 cN f2 l2
| _ -> raise ex)
| _, App (f2, l2) when flags.modulo_eta &&
- is_eta_constructor_app curenv flags.modulo_delta f2 l2 cM ->
+ is_eta_constructor_app curenv sigma flags.modulo_delta f2 l2 cM ->
(try
- let l2', l1' = eta_constructor_app curenv f2 l2 cM in
+ let l2', l1' = eta_constructor_app curenv sigma f2 l2 cM in
let opt' = {opt with at_top = true; with_cs = false} in
Array.fold_left2 (unirec_rec curenvnb CONV opt') substn l1' l2'
with ex when precatchable_exception ex ->
- match kind_of_term cM with
+ match EConstr.kind sigma cM with
| App(f1,l1) when
- (isMeta f1 && use_metas_pattern_unification flags nb l1
- || use_evars_pattern_unification flags && isAllowedEvar flags f1) ->
+ (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) ->
unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN f2 l2
| _ -> raise ex)
@@ -817,13 +858,13 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
reduce curenvnb pb opt substn cM cN)
| App (f1,l1), _ when
- (isMeta f1 && use_metas_pattern_unification flags nb l1
- || use_evars_pattern_unification flags && isAllowedEvar flags f1) ->
+ (isMeta sigma f1 && use_metas_pattern_unification sigma flags nb l1
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f1) ->
unify_app_pattern true curenvnb pb opt substn cM f1 l1 cN cN [||]
| _, App (f2,l2) when
- (isMeta f2 && use_metas_pattern_unification flags nb l2
- || use_evars_pattern_unification flags && isAllowedEvar flags f2) ->
+ (isMeta sigma f2 && use_metas_pattern_unification sigma flags nb l2
+ || use_evars_pattern_unification flags && isAllowedEvar sigma flags f2) ->
unify_app_pattern false curenvnb pb opt substn cM cM [||] cN f2 l2
| App (f1,l1), App (f2,l2) ->
@@ -838,11 +879,11 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| _ ->
unify_not_same_head curenvnb pb opt substn cM cN
- and unify_app_pattern dir curenvnb pb opt substn cM f1 l1 cN f2 l2 =
+ and unify_app_pattern dir curenvnb pb opt (sigma, _, _ as substn) cM f1 l1 cN f2 l2 =
let f, l, t = if dir then f1, l1, cN else f2, l2, cM in
match is_unification_pattern curenvnb sigma f (Array.to_list l) t with
| None ->
- (match kind_of_term t with
+ (match EConstr.kind sigma t with
| App (f',l') ->
if dir then unify_app curenvnb pb opt substn cM f1 l1 t f' l'
else unify_app curenvnb pb opt substn t f' l' cN f2 l2
@@ -851,19 +892,19 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| Some l ->
solve_pattern_eqn_array curenvnb f l t substn
- and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn) cM f1 l1 cN f2 l2 =
+ and unify_app (curenv, nb as curenvnb) pb opt (sigma, metas, evars as substn : subst0) cM f1 l1 cN f2 l2 =
try
let needs_expansion p c' =
- match kind_of_term c' with
+ match EConstr.kind sigma c' with
| Meta _ -> true
| Evar _ -> true
| Const (c, u) -> Constant.equal c (Projection.constant p)
| _ -> false
in
let expand_proj c c' l =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, t) when not (Projection.unfolded p) && needs_expansion p c' ->
- (try destApp (Retyping.expand_projection curenv sigma p t (Array.to_list l))
+ (try destApp sigma (Retyping.expand_projection curenv sigma p t (Array.to_list l))
with RetypeError _ -> (** Unification can be called on ill-typed terms, due
to FO and eta in particular, fail gracefully in that case *)
(c, l))
@@ -898,7 +939,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
ty1 ty2
with RetypeError _ -> substn
- and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn) cM cN =
+ and unify_not_same_head curenvnb pb opt (sigma, metas, evars as substn : subst0) cM cN =
try canonical_projections curenvnb pb opt cM cN substn
with ex when precatchable_exception ex ->
let sigma', b = constr_cmp cv_pb sigma flags cM cN in
@@ -907,24 +948,24 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
try reduce curenvnb pb opt substn cM cN
with ex when precatchable_exception ex ->
let (f1,l1) =
- match kind_of_term cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in
+ match EConstr.kind sigma cM with App (f,l) -> (f,l) | _ -> (cM,[||]) in
let (f2,l2) =
- match kind_of_term cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in
+ match EConstr.kind sigma cN with App (f,l) -> (f,l) | _ -> (cN,[||]) in
expand curenvnb pb opt substn cM f1 l1 cN f2 l2
and reduce curenvnb pb opt (sigma, metas, evars as substn) cM cN =
if use_full_betaiota flags && not (subterm_restriction opt flags) then
let cM' = do_reduce flags.modulo_delta curenvnb sigma cM in
- if not (Term.eq_constr cM cM') then
+ if not (EConstr.eq_constr sigma cM cM') then
unirec_rec curenvnb pb opt substn cM' cN
else
let cN' = do_reduce flags.modulo_delta curenvnb sigma cN in
- if not (Term.eq_constr cN cN') then
+ if not (EConstr.eq_constr sigma cN cN') then
unirec_rec curenvnb pb opt substn cM cN'
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
else error_cannot_unify (fst curenvnb) sigma (cM,cN)
- and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn) cM f1 l1 cN f2 l2 =
+ and expand (curenv,_ as curenvnb) pb opt (sigma,metasubst,evarsubst as substn : subst0) cM f1 l1 cN f2 l2 =
let res =
(* Try full conversion on meta-free terms. *)
(* Back to 1995 (later on called trivial_unify in 2002), the
@@ -943,10 +984,10 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
| None -> None
| Some convflags ->
let subst = ((if flags.use_metas_eagerly_in_conv_on_closed_terms then metasubst else ms), (if flags.use_evars_eagerly_in_conv_on_closed_terms then evarsubst else es)) in
- match subst_defined_metas_evars subst cM with
+ match subst_defined_metas_evars sigma subst cM with
| None -> (* some undefined Metas in cM *) None
| Some m1 ->
- match subst_defined_metas_evars subst cN with
+ match subst_defined_metas_evars sigma subst cN with
| None -> (* some undefined Metas in cN *) None
| Some n1 ->
(* No subterm restriction there, too much incompatibilities *)
@@ -970,7 +1011,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
match res with
| Some substn -> substn
| None ->
- let cf1 = key_of curenv opt flags f1 and cf2 = key_of curenv opt flags f2 in
+ let cf1 = key_of curenv sigma opt flags f1 and cf2 = key_of curenv sigma opt flags f2 in
match oracle_order curenv cf1 cf2 with
| None -> error_cannot_unify curenv sigma (cM,cN)
| Some true ->
@@ -1000,7 +1041,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
and canonical_projections (curenv, _ as curenvnb) pb opt cM cN (sigma,_,_ as substn) =
let f1 () =
- if isApp cM then
+ if isApp sigma cM then
let f1l1 = whd_nored_state sigma (cM,Stack.empty) in
if is_open_canonical_projection curenv sigma f1l1 then
let f2l2 = whd_nored_state sigma (cN,Stack.empty) in
@@ -1016,7 +1057,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
error_cannot_unify (fst curenvnb) sigma (cM,cN)
else
try f1 () with e when precatchable_exception e ->
- if isApp cN then
+ if isApp sigma cN then
let f2l2 = whd_nored_state sigma (cN, Stack.empty) in
if is_open_canonical_projection curenv sigma f2l2 then
let f1l1 = whd_nored_state sigma (cM, Stack.empty) in
@@ -1038,7 +1079,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
(evd,t2::ks, m-1)
else
let mv = new_meta () in
- let evd' = meta_declare mv (substl ks b) evd in
+ let evd' = meta_declare mv (EConstr.Unsafe.to_constr (substl ks b)) evd in
(evd', mkMeta mv :: ks, m - 1))
(sigma,[],List.length bs) bs
in
@@ -1050,7 +1091,7 @@ let rec unify_0_with_initial_metas (sigma,ms,es as subst) conv_at_top env cv_pb
let (substn,_,_) = Reductionops.Stack.fold2
(fun s u1 u -> unirec_rec curenvnb pb opt' s u1 (substl ks u))
substn params1 params in
- let (substn,_,_) = Reductionops.Stack.fold2 (unirec_rec curenvnb pb opt') substn ts ts1 in
+ let (substn,_,_) = Reductionops.Stack.fold2 (fun s u1 u2 -> unirec_rec curenvnb pb opt' s u1 u2) substn ts ts1 in
let app = mkApp (c, Array.rev_of_list ks) in
(* let substn = unirec_rec curenvnb pb b false substn t cN in *)
unirec_rec curenvnb pb opt' substn c1 app
@@ -1097,7 +1138,7 @@ let right = false
let rec unify_with_eta keptside flags env sigma c1 c2 =
(* Question: try whd_all on ci if not two lambdas? *)
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind sigma c1, EConstr.kind sigma c2 with
| (Lambda (na,t1,c1'), Lambda (_,t2,c2')) ->
let env' = push_rel_assum (na,t1) env in
let sigma,metas,evars = unify_0 env sigma CONV flags t1 t2 in
@@ -1206,7 +1247,8 @@ let applyHead env (type r) (evd : r Sigma.t) n c =
if Int.equal n 0 then
Sigma (c, evd, p)
else
- match kind_of_term (whd_all env (Sigma.to_evar_map evd) cty) with
+ let sigma = Sigma.to_evar_map evd in
+ match EConstr.kind sigma (whd_all env sigma cty) with
| Prod (_,c1,c2) ->
let Sigma (evar, evd', q) = Evarutil.new_evar env evd ~src:(Loc.ghost,Evar_kinds.GoalEvar) c1 in
apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) (p +> q) evd'
@@ -1214,8 +1256,8 @@ let applyHead env (type r) (evd : r Sigma.t) n c =
in
apprec n c (Typing.unsafe_type_of env (Sigma.to_evar_map evd) c) Sigma.refl evd
-let is_mimick_head ts f =
- match kind_of_term f with
+let is_mimick_head sigma ts f =
+ match EConstr.kind sigma f with
| Const (c,u) -> not (CClosure.is_transparent_constant ts c)
| Var id -> not (CClosure.is_transparent_variable ts id)
| (Rel _|Construct _|Ind _) -> true
@@ -1225,7 +1267,7 @@ let try_to_coerce env evd c cty tycon =
let j = make_judge c cty in
let (evd',j') = inh_conv_coerce_rigid_to true Loc.ghost env evd j tycon in
let evd' = Evarconv.solve_unif_constraints_with_heuristics env evd' in
- let evd' = Evd.map_metas_fvalue (nf_evar evd') evd' in
+ let evd' = Evd.map_metas_fvalue (fun c -> EConstr.Unsafe.to_constr (nf_evar evd' (EConstr.of_constr c))) evd' in
(evd',j'.uj_val)
let w_coerce_to_type env evd c cty mvty =
@@ -1284,25 +1326,25 @@ let solve_simple_evar_eqn ts env evd ev rhs =
or in evars, possibly generating new unification problems; if [b]
is true, unification of types of metas is required *)
-let w_merge env with_types flags (evd,metas,evars) =
+let w_merge env with_types flags (evd,metas,evars : subst0) =
let rec w_merge_rec evd metas evars eqns =
(* Process evars *)
match evars with
| (curenv,(evk,_ as ev),rhs)::evars' ->
if Evd.is_defined evd evk then
- let v = Evd.existential_value evd ev in
+ let v = mkEvar ev in
let (evd,metas',evars'') =
unify_0 curenv evd CONV flags rhs v in
w_merge_rec evd (metas'@metas) (evars''@evars') eqns
else begin
(* This can make rhs' ill-typed if metas are *)
- let rhs' = subst_meta_instances metas rhs in
- match kind_of_term rhs with
- | App (f,cl) when occur_meta rhs' ->
- if occur_evar evk rhs' then
+ let rhs' = subst_meta_instances evd metas rhs in
+ match EConstr.kind evd rhs with
+ | App (f,cl) when occur_meta evd rhs' ->
+ if occur_evar evd evk rhs' then
error_occur_check curenv evd evk rhs';
- if is_mimick_head flags.modulo_delta f then
+ if is_mimick_head evd flags.modulo_delta f then
let evd' =
mimick_undefined_evar evd flags f (Array.length cl) evk in
w_merge_rec evd' metas evars eqns
@@ -1342,20 +1384,20 @@ let w_merge env with_types flags (evd,metas,evars) =
if meta_defined evd mv then
let {rebus=c'},(status',_) = meta_fvalue evd mv in
let (take_left,st,(evd,metas',evars')) =
- merge_instances env evd flags status' status c' c
+ merge_instances env evd flags status' status (EConstr.of_constr c') c
in
let evd' =
if take_left then evd
- else meta_reassign mv (c,(st,TypeProcessed)) evd
+ else meta_reassign mv (EConstr.Unsafe.to_constr c,(st,TypeProcessed)) evd
in
w_merge_rec evd' (metas'@metas@metas'') (evars'@evars'') eqns
else
let evd' =
if occur_meta_evd evd mv c then
- if isMetaOf mv (whd_all env evd c) then evd
+ if isMetaOf evd mv (whd_all env evd c) then evd
else error_cannot_unify env evd (mkMeta mv,c)
else
- meta_assign mv (c,(status,TypeProcessed)) evd in
+ meta_assign mv (EConstr.Unsafe.to_constr c,(status,TypeProcessed)) evd in
w_merge_rec evd' (metas''@metas) evars'' eqns
| [] ->
(* Process type eqns *)
@@ -1381,17 +1423,17 @@ let w_merge env with_types flags (evd,metas,evars) =
let evd' = Sigma.to_evar_map evd' in
let (evd'',mc,ec) =
unify_0 sp_env evd' CUMUL flags
- (get_type_of sp_env evd' c) ev.evar_concl in
+ (get_type_of sp_env evd' c) (EConstr.of_constr ev.evar_concl) in
let evd''' = w_merge_rec evd'' mc ec [] in
if evd' == evd'''
- then Evd.define sp c evd'''
- else Evd.define sp (Evarutil.nf_evar evd''' c) evd''' in
+ then Evd.define sp (EConstr.Unsafe.to_constr c) evd'''
+ else Evd.define sp (EConstr.Unsafe.to_constr (Evarutil.nf_evar evd''' c)) evd''' in
let check_types evd =
let metas = Evd.meta_list evd in
let eqns = List.fold_left (fun acc (mv, b) ->
match b with
- | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, t.rebus) :: acc
+ | Clval (n, (t, (c, TypeNotProcessed)), v) -> (mv, c, EConstr.of_constr t.rebus) :: acc
| _ -> acc) [] metas
in w_merge_rec evd [] [] eqns
in
@@ -1402,6 +1444,11 @@ let w_merge env with_types flags (evd,metas,evars) =
in
if with_types then check_types res else res
+let retract_coercible_metas evd =
+ let (metas, evd) = retract_coercible_metas evd in
+ let map (mv, c, st) = (mv, EConstr.of_constr c, st) in
+ (List.map map metas, evd)
+
let w_unify_meta_types env ?(flags=default_unify_flags ()) evd =
let metas,evd = retract_coercible_metas evd in
w_merge env true flags.merge_unify_flags (evd,metas,[])
@@ -1419,13 +1466,17 @@ let w_unify_meta_types env ?(flags=default_unify_flags ()) evd =
let head_app sigma m =
fst (whd_nored_state sigma (m, Stack.empty))
+let isEvar_or_Meta sigma c = match EConstr.kind sigma c with
+| Evar _ | Meta _ -> true
+| _ -> false
+
let check_types env flags (sigma,_,_ as subst) m n =
- if isEvar_or_Meta (head_app sigma m) then
+ if isEvar_or_Meta sigma (head_app sigma m) then
unify_0_with_initial_metas subst true env CUMUL
flags
(get_type_of env sigma n)
(get_type_of env sigma m)
- else if isEvar_or_Meta (head_app sigma n) then
+ else if isEvar_or_Meta sigma (head_app sigma n) then
unify_0_with_initial_metas subst true env CUMUL
flags
(get_type_of env sigma m)
@@ -1476,22 +1527,22 @@ let iter_fail f a =
(* make_abstraction: a variant of w_unify_to_subterm which works on
contexts, with evars, and possibly with occurrences *)
-let indirectly_dependent c d decls =
- not (isVar c) &&
+let indirectly_dependent sigma c d decls =
+ not (isVar sigma c) &&
(* This test is not needed if the original term is a variable, but
it is needed otherwise, as e.g. when abstracting over "2" in
"forall H:0=2, H=H:>(0=1+1) -> 0=2." where there is now obvious
way to see that the second hypothesis depends indirectly over 2 *)
- List.exists (fun d' -> dependent_in_decl (mkVar (NamedDecl.get_id d')) d) decls
+ List.exists (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) decls
-let indirect_dependency d decls =
- decls |> List.filter (fun d' -> dependent_in_decl (mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id
+let indirect_dependency sigma d decls =
+ decls |> List.filter (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id
let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) =
let current_sigma = Sigma.to_evar_map current_sigma in
let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in
let sigma, subst = nf_univ_variables sigma in
- Sigma.Unsafe.of_pair (subst_univs_constr subst (nf_evar sigma c), sigma)
+ Sigma.Unsafe.of_pair (EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr (nf_evar sigma c))), sigma)
let default_matching_core_flags sigma =
let ts = Names.full_transparent_state in {
@@ -1543,17 +1594,17 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
modulo_conv_on_closed_terms = Some Names.full_transparent_state;
restrict_conv_on_strict_subterms = true } }
else default_matching_flags pending in
- let n = List.length (snd (decompose_app c)) in
+ let n = Array.length (snd (decompose_app_vect sigma c)) in
let matching_fun _ t =
try
let t',l2 =
if from_prefix_of_ind then
(* We check for fully applied subterms of the form "u u1 .. un" *)
(* of inductive type knowing only a prefix "u u1 .. ui" *)
- let t,l = decompose_app t in
+ let t,l = decompose_app sigma t in
let l1,l2 =
try List.chop n l with Failure _ -> raise (NotUnifiable None) in
- if not (List.for_all closed0 l2) then raise (NotUnifiable None)
+ if not (List.for_all (fun c -> Vars.closed0 sigma c) l2) then raise (NotUnifiable None)
else
applist (t,l1), l2
else t, [] in
@@ -1579,9 +1630,9 @@ let make_pattern_test from_prefix_of_ind is_correct_type env sigma (pending,c) =
(fun test -> match test.testing_state with
| None -> None
| Some (sigma,_,l) ->
- let c = applist (nf_evar sigma (local_strong whd_meta sigma c),l) in
+ let c = applist (nf_evar sigma (local_strong whd_meta sigma c), l) in
let univs, subst = nf_univ_variables sigma in
- Some (sigma,subst_univs_constr subst c))
+ Some (sigma,EConstr.of_constr (CVars.subst_univs_constr subst (EConstr.Unsafe.to_constr c))))
let make_eq_test env evd c =
let out cstr =
@@ -1592,7 +1643,7 @@ let make_eq_test env evd c =
let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
let id =
let t = match ty with Some t -> t | None -> get_type_of env sigma c in
- let x = id_of_name_using_hdchar (Global.env()) t name in
+ let x = id_of_name_using_hdchar (Global.env()) sigma t name in
let ids = ids_of_named_context (named_context env) in
if name == Anonymous then next_ident_away_in_goal x ids else
if mem_named_context_val x (named_context_val env) then
@@ -1602,17 +1653,18 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
x
in
let likefirst = clause_with_generic_occurrences occs in
- let mkvarid () = mkVar id in
+ let mkvarid () = EConstr.mkVar id in
let compute_dependency _ d (sign,depdecls) =
+ let d = map_named_decl EConstr.of_constr d in
let hyp = NamedDecl.get_id d in
match occurrences_of_hyp hyp occs with
| NoOccurrences, InHyp ->
(push_named_context_val d sign,depdecls)
| AllOccurrences, InHyp as occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
- let newdecl = replace_term_occ_decl_modulo occ test mkvarid d in
- if Context.Named.Declaration.equal d newdecl
- && not (indirectly_dependent c d depdecls)
+ let newdecl = replace_term_occ_decl_modulo sigma occ test mkvarid d in
+ if Context.Named.Declaration.equal (EConstr.eq_constr sigma) d newdecl
+ && not (indirectly_dependent sigma c d depdecls)
then
if check_occs && not (in_every_hyp occs)
then raise (PretypeError (env,sigma,NoOccurrenceFound (c,Some hyp)))
@@ -1621,7 +1673,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
(push_named_context_val newdecl sign, newdecl :: depdecls)
| occ ->
(* There are specific occurrences, hence not like first *)
- let newdecl = replace_term_occ_decl_modulo (AtOccs occ) test mkvarid d in
+ let newdecl = replace_term_occ_decl_modulo sigma (AtOccs occ) test mkvarid d in
(push_named_context_val newdecl sign, newdecl :: depdecls) in
try
let sign,depdecls =
@@ -1631,7 +1683,7 @@ let make_abstraction_core name (test,out) env sigma c ty occs check_occs concl =
| NoOccurrences -> concl
| occ ->
let occ = if likefirst then LikeFirst else AtOccs occ in
- replace_term_occ_modulo occ test mkvarid concl
+ replace_term_occ_modulo sigma occ test mkvarid concl
in
let lastlhyp =
if List.is_empty depdecls then None else Some (NamedDecl.get_id (List.last depdecls)) in
@@ -1663,7 +1715,7 @@ type abstraction_request =
type 'r abstraction_result =
Names.Id.t * named_context_val *
- Context.Named.Declaration.t list * Names.Id.t option *
+ named_declaration list * Names.Id.t option *
types * (constr, 'r) Sigma.sigma option
let make_abstraction env evd ccl abs =
@@ -1685,7 +1737,7 @@ let keyed_unify env evd kop =
| None -> fun _ -> true
| Some kop ->
fun cl ->
- let kc = Keys.constr_key cl in
+ let kc = Keys.constr_key (fun c -> EConstr.kind evd c) cl in
match kc with
| None -> false
| Some kc -> Keys.equiv_keys kop kc
@@ -1695,22 +1747,22 @@ let keyed_unify env evd kop =
Fails if no match is found *)
let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
let bestexn = ref None in
- let kop = Keys.constr_key op in
+ let kop = Keys.constr_key (fun c -> EConstr.kind evd c) op in
let rec matchrec cl =
- let cl = strip_outer_cast cl in
+ let cl = strip_outer_cast evd cl in
(try
- if closed0 cl && not (isEvar cl) && keyed_unify env evd kop cl then
+ if closed0 evd cl && not (isEvar evd cl) && keyed_unify env evd kop cl then
(try
if !keyed_unification then
- let f1, l1 = decompose_app_vect op in
- let f2, l2 = decompose_app_vect cl in
+ let f1, l1 = decompose_app_vect evd op in
+ let f2, l2 = decompose_app_vect evd cl in
w_typed_unify_array env evd flags f1 l1 f2 l2,cl
else w_typed_unify env evd CONV flags op cl,cl
with ex when Pretype_errors.unsatisfiable_exception ex ->
bestexn := Some ex; error "Unsat")
else error "Bound 1"
with ex when precatchable_exception ex ->
- (match kind_of_term cl with
+ (match EConstr.kind evd cl with
| App (f,args) ->
let n = Array.length args in
assert (n>0);
@@ -1771,7 +1823,7 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) =
let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
let return a b =
let (evd,c as a) = a () in
- if List.exists (fun (evd',c') -> Term.eq_constr c c') b then b else a :: b
+ if List.exists (fun (evd',c') -> EConstr.eq_constr evd' c c') b then b else a :: b
in
let fail str _ = error str in
let bind f g a =
@@ -1790,12 +1842,12 @@ let w_unify_to_subterm_all env evd ?(flags=default_unify_flags ()) (op,cl) =
in ffail 0
in
let rec matchrec cl =
- let cl = strip_outer_cast cl in
+ let cl = strip_outer_cast evd cl in
(bind
- (if closed0 cl
+ (if closed0 evd cl
then return (fun () -> w_typed_unify env evd CONV flags op cl,cl)
else fail "Bound 1")
- (match kind_of_term cl with
+ (match EConstr.kind evd cl with
| App (f,args) ->
let n = Array.length args in
assert (n>0);
@@ -1835,13 +1887,13 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
List.fold_right
(fun op (evd,l) ->
let op = whd_meta evd op in
- if isMeta op then
+ if isMeta evd op then
if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l)
- else error_abstraction_over_meta env evd hdmeta (destMeta op)
+ else error_abstraction_over_meta env evd hdmeta (destMeta evd op)
else
let allow_K = flags.allow_K_in_toplevel_higher_order_unification in
let flags =
- if occur_meta_or_existential op || !keyed_unification then
+ if unsafe_occur_meta_or_existential op || !keyed_unification then
(* This is up to delta for subterms w/o metas ... *)
flags
else
@@ -1850,7 +1902,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
unify pre-existing non frozen evars of the goal or of the
pattern *)
set_no_delta_flags flags in
- let t' = (strip_outer_cast op,t) in
+ let t' = (strip_outer_cast evd op,t) in
let (evd',cl) =
try
if is_keyed_unification () then
@@ -1866,11 +1918,11 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t =
(* w_unify_to_subterm does not go through evars, so
the next step, which was already in <= 8.4, is
needed at least for compatibility of rewrite *)
- dependent op t -> (evd,op)
+ dependent evd op t -> (evd,op)
in
if not allow_K &&
(* ensure we found a different instance *)
- List.exists (fun op -> Term.eq_constr op cl) l
+ List.exists (fun op -> EConstr.eq_constr evd' op cl) l
then error_non_linear_unification env evd hdmeta cl
else (evd',cl::l))
oplist
@@ -1902,10 +1954,10 @@ let secondOrderAbstractionAlgo dep =
let w_unify2 env evd flags dep cv_pb ty1 ty2 =
let c1, oplist1 = whd_nored_stack evd ty1 in
let c2, oplist2 = whd_nored_stack evd ty2 in
- match kind_of_term c1, kind_of_term c2 with
+ match EConstr.kind evd c1, EConstr.kind evd c2 with
| Meta p1, _ ->
(* Find the predicate *)
- secondOrderAbstractionAlgo dep env evd flags ty2 (p1,oplist1)
+ secondOrderAbstractionAlgo dep env evd flags ty2 (p1, oplist1)
| _, Meta p2 ->
(* Find the predicate *)
secondOrderAbstractionAlgo dep env evd flags ty1 (p2, oplist2)
@@ -1932,11 +1984,11 @@ let w_unify2 env evd flags dep cv_pb ty1 ty2 =
convertible and first-order otherwise. But if failed if e.g. the type of
Meta(1) had meta-variables in it. *)
let w_unify env evd cv_pb ?(flags=default_unify_flags ()) ty1 ty2 =
- let hd1,l1 = decompose_appvect (whd_nored evd ty1) in
- let hd2,l2 = decompose_appvect (whd_nored evd ty2) in
+ let hd1,l1 = decompose_app_vect evd (whd_nored evd ty1) in
+ let hd2,l2 = decompose_app_vect evd (whd_nored evd ty2) in
let is_empty1 = Array.is_empty l1 in
let is_empty2 = Array.is_empty l2 in
- match kind_of_term hd1, not is_empty1, kind_of_term hd2, not is_empty2 with
+ match EConstr.kind evd hd1, not is_empty1, EConstr.kind evd hd2, not is_empty2 with
(* Pattern case *)
| (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
when Int.equal (Array.length l1) (Array.length l2) ->
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
index c63502eae..148178f2f 100644
--- a/pretyping/unification.mli
+++ b/pretyping/unification.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Environ
open Evd
@@ -79,7 +80,7 @@ val finish_evar_resolution : ?flags:Pretyping.inference_flags ->
type 'r abstraction_result =
Names.Id.t * named_context_val *
- Context.Named.Declaration.t list * Names.Id.t option *
+ named_declaration list * Names.Id.t option *
types * (constr, 'r) Sigma.sigma option
val make_abstraction : env -> 'r Sigma.t -> constr ->
@@ -97,28 +98,29 @@ val abstract_list_all :
(* For tracing *)
-val w_merge : env -> bool -> core_unify_flags -> evar_map *
- (metavariable * constr * (instance_constraint * instance_typing_status)) list *
- (env * types pexistential * types) list -> evar_map
+type metabinding = (metavariable * constr * (instance_constraint * instance_typing_status))
+
+type subst0 =
+ (evar_map *
+ metabinding list *
+ (Environ.env * existential * t) list)
+
+val w_merge : env -> bool -> core_unify_flags -> subst0 -> evar_map
val unify_0 : Environ.env ->
Evd.evar_map ->
Evd.conv_pb ->
core_unify_flags ->
- Term.types ->
- Term.types ->
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list
+ types ->
+ types ->
+ subst0
val unify_0_with_initial_metas :
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list ->
+ subst0 ->
bool ->
Environ.env ->
Evd.conv_pb ->
core_unify_flags ->
- Term.types ->
- Term.types ->
- Evd.evar_map * Evd.metabinding list *
- (Environ.env * Term.types Term.pexistential * Term.constr) list
-
+ types ->
+ types ->
+ subst0
diff --git a/pretyping/vnorm.ml b/pretyping/vnorm.ml
index 75159bf8b..74998349b 100644
--- a/pretyping/vnorm.ml
+++ b/pretyping/vnorm.ml
@@ -99,13 +99,15 @@ let construct_of_constr_block = construct_of_constr false
let type_of_ind env (ind, u) =
type_of_inductive env (Inductive.lookup_mind_specif env ind, u)
-let build_branches_type env (mind,_ as _ind) mib mip u params dep p =
+let build_branches_type env sigma (mind,_ as _ind) mib mip u params dep p =
let rtbl = mip.mind_reloc_tbl in
(* [build_one_branch i cty] construit le type de la ieme branche (commence
a 0) et les lambda correspondant aux realargs *)
let build_one_branch i cty =
let typi = type_constructor mind mib u cty params in
- let decl,indapp = Reductionops.splay_prod env Evd.empty typi in
+ let decl,indapp = Reductionops.splay_prod env sigma (EConstr.of_constr typi) in
+ let decl = List.map (on_snd EConstr.Unsafe.to_constr) decl in
+ let indapp = EConstr.Unsafe.to_constr indapp in
let decl_with_letin,_ = decompose_prod_assum typi in
let ((ind,u),cargs) = find_rectype_a env indapp in
let nparams = Array.length params in
@@ -131,28 +133,28 @@ let build_case_type dep p realargs c =
(* La fonction de normalisation *)
-let rec nf_val env v t = nf_whd env (whd_val v) t
+let rec nf_val env sigma v t = nf_whd env sigma (whd_val v) t
-and nf_vtype env v = nf_val env v crazy_type
+and nf_vtype env sigma v = nf_val env sigma v crazy_type
-and nf_whd env whd typ =
+and nf_whd env sigma whd typ =
match whd with
| Vsort s -> mkSort s
| Vprod p ->
- let dom = nf_vtype env (dom p) in
+ let dom = nf_vtype env sigma (dom p) in
let name = Name (Id.of_string "x") in
let vc = body_of_vfun (nb_rel env) (codom p) in
- let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) vc in
+ let codom = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vc in
mkProd(name,dom,codom)
- | Vfun f -> nf_fun env f typ
- | Vfix(f,None) -> nf_fix env f
- | Vfix(f,Some vargs) -> fst (nf_fix_app env f vargs)
- | Vcofix(cf,_,None) -> nf_cofix env cf
+ | Vfun f -> nf_fun env sigma f typ
+ | Vfix(f,None) -> nf_fix env sigma f
+ | Vfix(f,Some vargs) -> fst (nf_fix_app env sigma f vargs)
+ | Vcofix(cf,_,None) -> nf_cofix env sigma cf
| Vcofix(cf,_,Some vargs) ->
- let cfd = nf_cofix env cf in
+ let cfd = nf_cofix env sigma cf in
let i,(_,ta,_) = destCoFix cfd in
let t = ta.(i) in
- let _, args = nf_args env vargs t in
+ let _, args = nf_args env sigma vargs t in
mkApp(cfd,args)
| Vconstr_const n ->
construct_of_constr_const env n typ
@@ -165,10 +167,10 @@ and nf_whd env whd typ =
| _ -> assert false
else (tag, 0) in
let capp,ctyp = construct_of_constr_block env tag typ in
- let args = nf_bargs env b ofs ctyp in
+ let args = nf_bargs env sigma b ofs ctyp in
mkApp(capp,args)
| Vatom_stk(Aid idkey, stk) ->
- constr_type_of_idkey env idkey stk
+ constr_type_of_idkey env sigma idkey stk
| Vatom_stk(Aind ((mi,i) as ind), stk) ->
let mib = Environ.lookup_mind mi env in
let nb_univs =
@@ -178,12 +180,12 @@ and nf_whd env whd typ =
let mk u =
let pind = (ind, u) in (mkIndU pind, type_of_ind env pind)
in
- nf_univ_args ~nb_univs mk env stk
+ nf_univ_args ~nb_univs mk env sigma stk
| Vatom_stk(Atype u, stk) -> assert false
| Vuniv_level lvl ->
assert false
-and nf_univ_args ~nb_univs mk env stk =
+and nf_univ_args ~nb_univs mk env sigma stk =
let u =
if Int.equal nb_univs 0 then Univ.Instance.empty
else match stk with
@@ -195,9 +197,9 @@ and nf_univ_args ~nb_univs mk env stk =
| _ -> assert false
in
let (t,ty) = mk u in
- nf_stk ~from:nb_univs env t ty stk
+ nf_stk ~from:nb_univs env sigma t ty stk
-and constr_type_of_idkey env (idkey : Vars.id_key) stk =
+and constr_type_of_idkey env sigma (idkey : Vars.id_key) stk =
match idkey with
| ConstKey cst ->
let cbody = Environ.lookup_constant cst env in
@@ -208,30 +210,30 @@ and constr_type_of_idkey env (idkey : Vars.id_key) stk =
let mk u =
let pcst = (cst, u) in (mkConstU pcst, Typeops.type_of_constant_in env pcst)
in
- nf_univ_args ~nb_univs mk env stk
+ nf_univ_args ~nb_univs mk env sigma stk
| VarKey id ->
let ty = NamedDecl.get_type (lookup_named id env) in
- nf_stk env (mkVar id) ty stk
+ nf_stk env sigma (mkVar id) ty stk
| RelKey i ->
let n = (nb_rel env - i) in
let ty = RelDecl.get_type (lookup_rel n env) in
- nf_stk env (mkRel n) (lift n ty) stk
+ nf_stk env sigma (mkRel n) (lift n ty) stk
-and nf_stk ?from:(from=0) env c t stk =
+and nf_stk ?from:(from=0) env sigma c t stk =
match stk with
| [] -> c
| Zapp vargs :: stk ->
if nargs vargs >= from then
- let t, args = nf_args ~from:from env vargs t in
- nf_stk env (mkApp(c,args)) t stk
+ let t, args = nf_args ~from:from env sigma vargs t in
+ nf_stk env sigma (mkApp(c,args)) t stk
else
let rest = from - nargs vargs in
- nf_stk ~from:rest env c t stk
+ nf_stk ~from:rest env sigma c t stk
| Zfix (f,vargs) :: stk ->
assert (from = 0) ;
- let fa, typ = nf_fix_app env f vargs in
+ let fa, typ = nf_fix_app env sigma f vargs in
let _,_,codom = decompose_prod env typ in
- nf_stk env (mkApp(fa,[|c|])) (subst1 c codom) stk
+ nf_stk env sigma (mkApp(fa,[|c|])) (subst1 c codom) stk
| Zswitch sw :: stk ->
assert (from = 0) ;
let ((mind,_ as ind), u), allargs = find_rectype_a env t in
@@ -241,34 +243,34 @@ and nf_stk ?from:(from=0) env c t stk =
let pT =
hnf_prod_applist env (type_of_ind env (ind,u)) (Array.to_list params) in
let pT = whd_all env pT in
- let dep, p = nf_predicate env (ind,u) mip params (type_of_switch sw) pT in
+ let dep, p = nf_predicate env sigma (ind,u) mip params (type_of_switch sw) pT in
(* Calcul du type des branches *)
- let btypes = build_branches_type env ind mib mip u params dep p in
+ let btypes = build_branches_type env sigma ind mib mip u params dep p in
(* calcul des branches *)
let bsw = branch_of_switch (nb_rel env) sw in
let mkbranch i (n,v) =
let decl,decl_with_letin,codom = btypes.(i) in
- let b = nf_val (Termops.push_rels_assum decl env) v codom in
+ let b = nf_val (Termops.push_rels_assum decl env) sigma v codom in
Termops.it_mkLambda_or_LetIn_from_no_LetIn b decl_with_letin
in
let branchs = Array.mapi mkbranch bsw in
let tcase = build_case_type dep p realargs c in
let ci = case_info sw in
- nf_stk env (mkCase(ci, p, c, branchs)) tcase stk
+ nf_stk env sigma (mkCase(ci, p, c, branchs)) tcase stk
| Zproj p :: stk ->
assert (from = 0) ;
let p' = Projection.make p true in
- let ty = Inductiveops.type_of_projection_knowing_arg env Evd.empty p' c t in
- nf_stk env (mkProj(p',c)) ty stk
+ let ty = Inductiveops.type_of_projection_knowing_arg env sigma p' (EConstr.of_constr c) (EConstr.of_constr t) in
+ nf_stk env sigma (mkProj(p',c)) ty stk
-and nf_predicate env ind mip params v pT =
+and nf_predicate env sigma ind mip params v pT =
match whd_val v, kind_of_term pT with
| Vfun f, Prod _ ->
let k = nb_rel env in
let vb = body_of_vfun k f in
let name,dom,codom = decompose_prod env pT in
let dep,body =
- nf_predicate (push_rel (LocalAssum (name,dom)) env) ind mip params vb codom in
+ nf_predicate (push_rel (LocalAssum (name,dom)) env) sigma ind mip params vb codom in
dep, mkLambda(name,dom,body)
| Vfun f, _ ->
let k = nb_rel env in
@@ -278,33 +280,33 @@ and nf_predicate env ind mip params v pT =
let rargs = Array.init n (fun i -> mkRel (n-i)) in
let params = if Int.equal n 0 then params else Array.map (lift n) params in
let dom = mkApp(mkIndU ind,Array.append params rargs) in
- let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) vb in
+ let body = nf_vtype (push_rel (LocalAssum (name,dom)) env) sigma vb in
true, mkLambda(name,dom,body)
- | _, _ -> false, nf_val env v crazy_type
+ | _, _ -> false, nf_val env sigma v crazy_type
-and nf_args env vargs ?from:(f=0) t =
+and nf_args env sigma vargs ?from:(f=0) t =
let t = ref t in
let len = nargs vargs - f in
let args =
Array.init len
(fun i ->
let _,dom,codom = decompose_prod env !t in
- let c = nf_val env (arg vargs (f+i)) dom in
+ let c = nf_val env sigma (arg vargs (f+i)) dom in
t := subst1 c codom; c) in
!t,args
-and nf_bargs env b ofs t =
+and nf_bargs env sigma b ofs t =
let t = ref t in
let len = bsize b - ofs in
let args =
Array.init len
(fun i ->
let _,dom,codom = decompose_prod env !t in
- let c = nf_val env (bfield b (i+ofs)) dom in
+ let c = nf_val env sigma (bfield b (i+ofs)) dom in
t := subst1 c codom; c) in
args
-and nf_fun env f typ =
+and nf_fun env sigma f typ =
let k = nb_rel env in
let vb = body_of_vfun k f in
let name,dom,codom =
@@ -314,46 +316,51 @@ and nf_fun env f typ =
CErrors.anomaly
(Pp.strbrk "Returned a functional value in a type not recognized as a product type.")
in
- let body = nf_val (push_rel (LocalAssum (name,dom)) env) vb codom in
+ let body = nf_val (push_rel (LocalAssum (name,dom)) env) sigma vb codom in
mkLambda(name,dom,body)
-and nf_fix env f =
+and nf_fix env sigma f =
let init = current_fix f in
let rec_args = rec_args f in
let k = nb_rel env in
let vb, vt = reduce_fix k f in
let ndef = Array.length vt in
- let ft = Array.map (fun v -> nf_val env v crazy_type) vt in
+ let ft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (Id.of_string "Ffix"))) in
(* Third argument of the tuple is ignored by push_rec_types *)
let env = push_rec_types (name,ft,ft) env in
(* We lift here because the types of arguments (in tt) will be evaluated
in an environment where the fixpoints have been pushed *)
- let norm_vb v t = nf_fun env v (lift ndef t) in
+ let norm_vb v t = nf_fun env sigma v (lift ndef t) in
let fb = Util.Array.map2 norm_vb vb ft in
mkFix ((rec_args,init),(name,ft,fb))
-and nf_fix_app env f vargs =
- let fd = nf_fix env f in
+and nf_fix_app env sigma f vargs =
+ let fd = nf_fix env sigma f in
let (_,i),(_,ta,_) = destFix fd in
let t = ta.(i) in
- let t, args = nf_args env vargs t in
+ let t, args = nf_args env sigma vargs t in
mkApp(fd,args),t
-and nf_cofix env cf =
+and nf_cofix env sigma cf =
let init = current_cofix cf in
let k = nb_rel env in
let vb,vt = reduce_cofix k cf in
let ndef = Array.length vt in
- let cft = Array.map (fun v -> nf_val env v crazy_type) vt in
+ let cft = Array.map (fun v -> nf_val env sigma v crazy_type) vt in
let name = Array.init ndef (fun _ -> (Name (Id.of_string "Fcofix"))) in
let env = push_rec_types (name,cft,cft) env in
- let cfb = Util.Array.map2 (fun v t -> nf_val env v t) vb cft in
+ let cfb = Util.Array.map2 (fun v t -> nf_val env sigma v t) vb cft in
mkCoFix (init,(name,cft,cfb))
-let cbv_vm env c t =
+let cbv_vm env sigma c t =
+ if Termops.occur_meta_or_existential sigma c then
+ CErrors.error "vm_compute does not support existential variables.";
+ (** This evar-normalizes terms beforehand *)
+ let c = EConstr.to_constr sigma c in
+ let t = EConstr.to_constr sigma t in
let v = Vconv.val_of_constr env c in
- nf_val env v t
+ EConstr.of_constr (nf_val env sigma v t)
let vm_infer_conv ?(pb=Reduction.CUMUL) env sigma t1 t2 =
Reductionops.infer_conv_gen (fun pb ~l2r sigma ts -> Vconv.vm_conv_gen pb)
diff --git a/pretyping/vnorm.mli b/pretyping/vnorm.mli
index 58f5b14e1..8a4202c88 100644
--- a/pretyping/vnorm.mli
+++ b/pretyping/vnorm.mli
@@ -6,8 +6,8 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-open Term
+open EConstr
open Environ
(** {6 Reduction functions } *)
-val cbv_vm : env -> constr -> types -> constr
+val cbv_vm : env -> Evd.evar_map -> constr -> types -> constr
diff --git a/printing/prettyp.ml b/printing/prettyp.ml
index 5963d45ef..aa422e36a 100644
--- a/printing/prettyp.ml
+++ b/printing/prettyp.ml
@@ -42,8 +42,8 @@ type object_pr = {
print_named_decl : Context.Named.Declaration.t -> std_ppcmds;
print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option;
print_context : bool -> int option -> Lib.library_segment -> std_ppcmds;
- print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds;
- print_eval : Reductionops.reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds;
+ print_typed_value_in_env : Environ.env -> Evd.evar_map -> EConstr.constr * EConstr.types -> Pp.std_ppcmds;
+ print_eval : Reductionops.reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds;
}
let gallina_print_module = print_module
@@ -71,10 +71,11 @@ let print_basename sp = pr_global (ConstRef sp)
let print_ref reduce ref =
let typ = Global.type_of_global_unsafe ref in
+ let typ = EConstr.of_constr typ in
let typ =
if reduce then
let ctx,ccl = Reductionops.splay_prod_assum (Global.env()) Evd.empty typ
- in it_mkProd_or_LetIn ccl ctx
+ in EConstr.it_mkProd_or_LetIn ccl ctx
else typ in
let univs = Global.universes_of_global ref in
let env = Global.env () in
@@ -84,7 +85,7 @@ let print_ref reduce ref =
if Global.is_polymorphic ref then Printer.pr_universe_instance sigma univs
else mt ()
in
- hov 0 (pr_global ref ++ inst ++ str " :" ++ spc () ++ pr_ltype_env env sigma typ ++
+ hov 0 (pr_global ref ++ inst ++ str " :" ++ spc () ++ pr_letype_env env sigma typ ++
Printer.pr_universe_ctx sigma univs)
(********************************)
@@ -438,8 +439,8 @@ let print_located_qualid ref = print_located_qualid "object" [`TERM; `LTAC; `MOD
(**** Gallina layer *****)
let gallina_print_typed_value_in_env env sigma (trm,typ) =
- (pr_lconstr_env env sigma trm ++ fnl () ++
- str " : " ++ pr_ltype_env env sigma typ)
+ (pr_leconstr_env env sigma trm ++ fnl () ++
+ str " : " ++ pr_letype_env env sigma typ)
(* To be improved; the type should be used to provide the types in the
abstractions. This should be done recursively inside pr_lconstr, so that
@@ -647,6 +648,8 @@ let print_judgment env sigma {uj_val=trm;uj_type=typ} =
let print_safe_judgment env sigma j =
let trm = Safe_typing.j_val j in
let typ = Safe_typing.j_type j in
+ let trm = EConstr.of_constr trm in
+ let typ = EConstr.of_constr typ in
print_typed_value_in_env env sigma (trm, typ)
(*********************)
@@ -766,7 +769,9 @@ let print_opaque_name qid =
| IndRef (sp,_) ->
print_inductive sp
| ConstructRef cstr as gr ->
+ let open EConstr in
let ty = Universes.unsafe_type_of_global gr in
+ let ty = EConstr.of_constr ty in
print_typed_value (mkConstruct cstr, ty)
| VarRef id ->
env |> lookup_named id |> NamedDecl.set_id id |> print_named_decl
diff --git a/printing/prettyp.mli b/printing/prettyp.mli
index 0eab15579..38e111034 100644
--- a/printing/prettyp.mli
+++ b/printing/prettyp.mli
@@ -27,11 +27,11 @@ val print_full_context_typ : unit -> std_ppcmds
val print_full_pure_context : unit -> std_ppcmds
val print_sec_context : reference -> std_ppcmds
val print_sec_context_typ : reference -> std_ppcmds
-val print_judgment : env -> Evd.evar_map -> unsafe_judgment -> std_ppcmds
+val print_judgment : env -> Evd.evar_map -> EConstr.unsafe_judgment -> std_ppcmds
val print_safe_judgment : env -> Evd.evar_map -> Safe_typing.judgment -> std_ppcmds
val print_eval :
reduction_function -> env -> Evd.evar_map ->
- Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds
+ Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds
val print_name : reference or_by_notation -> std_ppcmds
val print_opaque_name : reference -> std_ppcmds
@@ -69,8 +69,8 @@ type object_pr = {
print_named_decl : Context.Named.Declaration.t -> std_ppcmds;
print_library_entry : bool -> (object_name * Lib.node) -> std_ppcmds option;
print_context : bool -> int option -> Lib.library_segment -> std_ppcmds;
- print_typed_value_in_env : Environ.env -> Evd.evar_map -> Term.constr * Term.types -> Pp.std_ppcmds;
- print_eval : reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> unsafe_judgment -> std_ppcmds
+ print_typed_value_in_env : Environ.env -> Evd.evar_map -> EConstr.constr * EConstr.types -> Pp.std_ppcmds;
+ print_eval : reduction_function -> env -> Evd.evar_map -> Constrexpr.constr_expr -> EConstr.unsafe_judgment -> std_ppcmds
}
val set_object_pr : object_pr -> unit
diff --git a/printing/printer.ml b/printing/printer.ml
index 5e7e9ce54..239e1d7eb 100644
--- a/printing/printer.ml
+++ b/printing/printer.ml
@@ -60,7 +60,10 @@ let pr_constr_goal_style_env env = pr_constr_core true env
let pr_open_lconstr_env env sigma (_,c) = pr_lconstr_env env sigma c
let pr_open_constr_env env sigma (_,c) = pr_constr_env env sigma c
- (* NB do not remove the eta-redexes! Global.env() has side-effects... *)
+let pr_leconstr_env env sigma c = pr_lconstr_env env sigma (EConstr.to_constr sigma c)
+let pr_econstr_env env sigma c = pr_constr_env env sigma (EConstr.to_constr sigma c)
+
+(* NB do not remove the eta-redexes! Global.env() has side-effects... *)
let pr_lconstr t =
let (sigma, env) = get_current_context () in
pr_lconstr_env env sigma t
@@ -71,6 +74,9 @@ let pr_constr t =
let pr_open_lconstr (_,c) = pr_lconstr c
let pr_open_constr (_,c) = pr_constr c
+let pr_leconstr c = pr_lconstr (EConstr.Unsafe.to_constr c)
+let pr_econstr c = pr_constr (EConstr.Unsafe.to_constr c)
+
let pr_constr_under_binders_env_gen pr env sigma (ids,c) =
(* Warning: clashes can occur with variables of same name in env but *)
(* we also need to preserve the actual names of the patterns *)
@@ -78,8 +84,8 @@ let pr_constr_under_binders_env_gen pr env sigma (ids,c) =
let assums = List.map (fun id -> (Name id,(* dummy *) mkProp)) ids in
pr (Termops.push_rels_assum assums env) sigma c
-let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_constr_env
-let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_lconstr_env
+let pr_constr_under_binders_env = pr_constr_under_binders_env_gen pr_econstr_env
+let pr_lconstr_under_binders_env = pr_constr_under_binders_env_gen pr_leconstr_env
let pr_constr_under_binders c =
let (sigma, env) = get_current_context () in
@@ -93,7 +99,6 @@ let pr_type_core goal_concl_style env sigma t =
let pr_ltype_core goal_concl_style env sigma t =
pr_lconstr_expr (extern_type goal_concl_style env sigma t)
-let pr_goal_concl_style_env env = pr_ltype_core true env
let pr_ltype_env env = pr_ltype_core false env
let pr_type_env env = pr_type_core false env
@@ -104,8 +109,13 @@ let pr_type t =
let (sigma, env) = get_current_context () in
pr_type_env env sigma t
+let pr_etype_env env sigma c = pr_type_env env sigma (EConstr.to_constr sigma c)
+let pr_letype_env env sigma c = pr_ltype_env env sigma (EConstr.to_constr sigma c)
+let pr_goal_concl_style_env env sigma c =
+ pr_ltype_core true env sigma (EConstr.to_constr sigma c)
+
let pr_ljudge_env env sigma j =
- (pr_lconstr_env env sigma j.uj_val, pr_lconstr_env env sigma j.uj_type)
+ (pr_leconstr_env env sigma j.uj_val, pr_leconstr_env env sigma j.uj_type)
let pr_ljudge j =
let (sigma, env) = get_current_context () in
@@ -147,7 +157,7 @@ let pr_constr_pattern t =
let pr_sort sigma s = pr_glob_sort (extern_sort sigma s)
let _ = Termops.set_print_constr
- (fun env t -> pr_lconstr_expr (extern_constr ~lax:true false env Evd.empty t))
+ (fun env sigma t -> pr_lconstr_expr (extern_constr ~lax:true false env sigma (EConstr.Unsafe.to_constr t)))
let pr_in_comment pr x = str "(* " ++ pr x ++ str " *)"
@@ -213,7 +223,7 @@ let safe_pr_constr t =
let pr_universe_ctx sigma c =
if !Detyping.print_universes && not (Univ.UContext.is_empty c) then
fnl()++pr_in_comment (fun c -> v 0
- (Univ.pr_universe_context (Evd.pr_evd_level sigma) c)) c
+ (Univ.pr_universe_context (Termops.pr_evd_level sigma) c)) c
else
mt()
@@ -230,7 +240,7 @@ let pr_puniverses f env (c,u) =
else mt ())
let pr_constant env cst = pr_global_env (Termops.vars_of_env env) (ConstRef cst)
-let pr_existential_key = Evd.pr_existential_key
+let pr_existential_key = Termops.pr_existential_key
let pr_existential env sigma ev = pr_lconstr_env env sigma (mkEvar ev)
let pr_inductive env ind = pr_lconstr_env env Evd.empty (mkInd ind)
let pr_constructor env cstr = pr_lconstr_env env Evd.empty (mkConstruct cstr)
@@ -496,15 +506,15 @@ let print_evar_constraints gl sigma =
| Some g ->
let env = Goal.V82.env sigma g in fun e' ->
begin
- if Context.Named.equal (named_context env) (named_context e') then
- if Context.Rel.equal (rel_context env) (rel_context e') then mt ()
+ if Context.Named.equal Constr.equal (named_context env) (named_context e') then
+ if Context.Rel.equal Constr.equal (rel_context env) (rel_context e') then mt ()
else pr_rel_context_of e' sigma ++ str " |-" ++ spc ()
else pr_context_of e' sigma ++ str " |-" ++ spc ()
end
in
let pr_evconstr (pbty,env,t1,t2) =
- let t1 = Evarutil.nf_evar sigma t1
- and t2 = Evarutil.nf_evar sigma t2 in
+ let t1 = Evarutil.nf_evar sigma (EConstr.of_constr t1)
+ and t2 = Evarutil.nf_evar sigma (EConstr.of_constr t2) in
let env =
(** We currently allow evar instances to refer to anonymous de Bruijn
indices, so we protect the error printing code in this case by giving
@@ -512,13 +522,13 @@ let print_evar_constraints gl sigma =
problem. MS: we should rather stop depending on anonymous variables, they
can be used to indicate independency. Also, this depends on a strategy for
naming/renaming *)
- Namegen.make_all_name_different env in
+ Namegen.make_all_name_different env sigma in
str" " ++
- hov 2 (pr_env env ++ pr_lconstr_env env sigma t1 ++ spc () ++
+ hov 2 (pr_env env ++ pr_leconstr_env env sigma t1 ++ spc () ++
str (match pbty with
| Reduction.CONV -> "=="
| Reduction.CUMUL -> "<=") ++
- spc () ++ pr_lconstr_env env sigma t2)
+ spc () ++ pr_leconstr_env env sigma t2)
in
let pr_candidate ev evi (candidates,acc) =
if Option.has_some evi.evar_candidates then
@@ -767,7 +777,8 @@ let pr_prim_rule = function
str ";[" ++ cl ++ str"intro " ++ pr_id id ++ str"|idtac]")
| Refine c ->
- str(if Termops.occur_meta c then "refine " else "exact ") ++
+ (** FIXME *)
+ str(if Termops.occur_meta Evd.empty (EConstr.of_constr c) then "refine " else "exact ") ++
Constrextern.with_meta_as_hole pr_constr c
(* Backwards compatibility *)
@@ -918,4 +929,4 @@ let pr_polymorphic b =
let pr_universe_instance evd ctx =
let inst = Univ.UContext.instance ctx in
- str"@{" ++ Univ.Instance.pr (Evd.pr_evd_level evd) inst ++ str"}"
+ str"@{" ++ Univ.Instance.pr (Termops.pr_evd_level evd) inst ++ str"}"
diff --git a/printing/printer.mli b/printing/printer.mli
index 20032012a..504392e35 100644
--- a/printing/printer.mli
+++ b/printing/printer.mli
@@ -38,6 +38,13 @@ val safe_pr_lconstr : constr -> std_ppcmds
val safe_pr_constr_env : env -> evar_map -> constr -> std_ppcmds
val safe_pr_constr : constr -> std_ppcmds
+val pr_econstr_env : env -> evar_map -> EConstr.t -> std_ppcmds
+val pr_econstr : EConstr.t -> std_ppcmds
+val pr_leconstr_env : env -> evar_map -> EConstr.t -> std_ppcmds
+val pr_leconstr : EConstr.t -> std_ppcmds
+
+val pr_etype_env : env -> evar_map -> EConstr.types -> std_ppcmds
+val pr_letype_env : env -> evar_map -> EConstr.types -> std_ppcmds
val pr_open_constr_env : env -> evar_map -> open_constr -> std_ppcmds
val pr_open_constr : open_constr -> std_ppcmds
@@ -51,7 +58,7 @@ val pr_constr_under_binders : constr_under_binders -> std_ppcmds
val pr_lconstr_under_binders_env : env -> evar_map -> constr_under_binders -> std_ppcmds
val pr_lconstr_under_binders : constr_under_binders -> std_ppcmds
-val pr_goal_concl_style_env : env -> evar_map -> types -> std_ppcmds
+val pr_goal_concl_style_env : env -> evar_map -> EConstr.types -> std_ppcmds
val pr_ltype_env : env -> evar_map -> types -> std_ppcmds
val pr_ltype : types -> std_ppcmds
@@ -61,8 +68,8 @@ val pr_type : types -> std_ppcmds
val pr_closed_glob_env : env -> evar_map -> closed_glob_constr -> std_ppcmds
val pr_closed_glob : closed_glob_constr -> std_ppcmds
-val pr_ljudge_env : env -> evar_map -> unsafe_judgment -> std_ppcmds * std_ppcmds
-val pr_ljudge : unsafe_judgment -> std_ppcmds * std_ppcmds
+val pr_ljudge_env : env -> evar_map -> EConstr.unsafe_judgment -> std_ppcmds * std_ppcmds
+val pr_ljudge : EConstr.unsafe_judgment -> std_ppcmds * std_ppcmds
val pr_lglob_constr_env : env -> glob_constr -> std_ppcmds
val pr_lglob_constr : glob_constr -> std_ppcmds
diff --git a/printing/printmod.ml b/printing/printmod.ml
index baa1b8d79..6f4b162d7 100644
--- a/printing/printmod.ml
+++ b/printing/printmod.ml
@@ -94,7 +94,7 @@ let print_one_inductive env sigma mib ((_,i) as ind) =
else Univ.Instance.empty in
let mip = mib.mind_packets.(i) in
let params = Inductive.inductive_paramdecls (mib,u) in
- let args = Context.Rel.to_extended_list 0 params in
+ let args = Context.Rel.to_extended_list mkRel 0 params in
let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in
let cstrtypes = Inductive.type_of_constructors (ind,u) (mib,mip) in
let cstrtypes = Array.map (fun c -> hnf_prod_applist env c args) cstrtypes in
@@ -148,7 +148,7 @@ let print_record env mind mib =
in
let mip = mib.mind_packets.(0) in
let params = Inductive.inductive_paramdecls (mib,u) in
- let args = Context.Rel.to_extended_list 0 params in
+ let args = Context.Rel.to_extended_list mkRel 0 params in
let arity = hnf_prod_applist env (build_ind_type env ((mib,mip),u)) args in
let cstrtypes = Inductive.type_of_constructors ((mind,0),u) (mib,mip) in
let cstrtype = hnf_prod_applist env cstrtypes.(0) args in
diff --git a/proofs/clenv.ml b/proofs/clenv.ml
index fad656223..f9ebc4233 100644
--- a/proofs/clenv.ml
+++ b/proofs/clenv.ml
@@ -12,11 +12,12 @@ open Util
open Names
open Nameops
open Term
-open Vars
open Termops
open Namegen
open Environ
open Evd
+open EConstr
+open Vars
open Reduction
open Reductionops
open Tacred
@@ -43,12 +44,6 @@ type clausenv = {
let cl_env ce = ce.env
let cl_sigma ce = ce.evd
-let map_clenv sub clenv =
- { templval = map_fl sub clenv.templval;
- templtyp = map_fl sub clenv.templtyp;
- evd = cmap sub clenv.evd;
- env = clenv.env }
-
let clenv_nf_meta clenv c = nf_meta clenv.evd c
let clenv_term clenv c = meta_instance clenv.evd c
let clenv_meta_type clenv mv = Typing.meta_type clenv.evd mv
@@ -56,9 +51,9 @@ let clenv_value clenv = meta_instance clenv.evd clenv.templval
let clenv_type clenv = meta_instance clenv.evd clenv.templtyp
let refresh_undefined_univs clenv =
- match kind_of_term clenv.templval.rebus with
+ match EConstr.kind clenv.evd clenv.templval.rebus with
| Var _ -> clenv, Univ.empty_level_subst
- | App (f, args) when isVar f -> clenv, Univ.empty_level_subst
+ | App (f, args) when isVar clenv.evd f -> clenv, Univ.empty_level_subst
| _ ->
let evd', subst = Evd.refresh_undefined_universes clenv.evd in
let map_freelisted f = { f with rebus = subst_univs_level_constr subst f.rebus } in
@@ -71,15 +66,18 @@ let clenv_get_type_of ce c = Retyping.get_type_of (cl_env ce) (cl_sigma ce) c
exception NotExtensibleClause
+let mk_freelisted c =
+ map_fl EConstr.of_constr (mk_freelisted (EConstr.Unsafe.to_constr c))
+
let clenv_push_prod cl =
let typ = whd_all (cl_env cl) (cl_sigma cl) (clenv_type cl) in
- let rec clrec typ = match kind_of_term typ with
+ let rec clrec typ = match EConstr.kind cl.evd typ with
| Cast (t,_,_) -> clrec t
| Prod (na,t,u) ->
let mv = new_meta () in
- let dep = dependent (mkRel 1) u in
+ let dep = not (noccurn (cl_sigma cl) 1 u) in
let na' = if dep then na else Anonymous in
- let e' = meta_declare mv t ~name:na' cl.evd in
+ let e' = meta_declare mv (EConstr.Unsafe.to_constr t) ~name:na' cl.evd in
let concl = if dep then subst1 (mkMeta mv) u else u in
let def = applist (cl.templval.rebus,[mkMeta mv]) in
{ templval = mk_freelisted def;
@@ -103,14 +101,17 @@ let clenv_push_prod cl =
and [ccl] is [forall y, Meta n1=y -> y=Meta n1] *)
let clenv_environments evd bound t =
+ let open EConstr in
+ let open Vars in
let rec clrec (e,metas) n t =
- match n, kind_of_term t with
+ match n, EConstr.kind evd t with
| (Some 0, _) -> (e, List.rev metas, t)
| (n, Cast (t,_,_)) -> clrec (e,metas) n t
| (n, Prod (na,t1,t2)) ->
let mv = new_meta () in
- let dep = not (noccurn 1 t2) in
+ let dep = not (noccurn evd 1 t2) in
let na' = if dep then na else Anonymous in
+ let t1 = EConstr.Unsafe.to_constr t1 in
let e' = meta_declare mv t1 ~name:na' e in
clrec (e', (mkMeta mv)::metas) (Option.map ((+) (-1)) n)
(if dep then (subst1 (mkMeta mv) t2) else t2)
@@ -128,11 +129,13 @@ let mk_clenv_from_env env sigma n (c,cty) =
env = env }
let mk_clenv_from_n gls n (c,cty) =
- mk_clenv_from_env (pf_env gls) gls.sigma n (c, cty)
+ let env = Proofview.Goal.env gls in
+ let sigma = Tacmach.New.project gls in
+ mk_clenv_from_env env sigma n (c, cty)
let mk_clenv_from gls = mk_clenv_from_n gls None
-let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t)
+let mk_clenv_type_of gls t = mk_clenv_from gls (t,Tacmach.New.pf_unsafe_type_of gls t)
(******************************************************************)
@@ -168,13 +171,13 @@ let clenv_assign mv rhs clenv =
error "clenv_assign: circularity in unification";
try
if meta_defined clenv.evd mv then
- if not (Term.eq_constr (fst (meta_fvalue clenv.evd mv)).rebus rhs) then
+ if not (EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd mv)).rebus) rhs) then
error_incompatible_inst clenv mv
else
clenv
else
let st = (Conv,TypeNotProcessed) in
- {clenv with evd = meta_assign mv (rhs_fls.rebus,st) clenv.evd}
+ {clenv with evd = meta_assign mv (EConstr.Unsafe.to_constr rhs_fls.rebus,st) clenv.evd}
with Not_found ->
error "clenv_assign: undefined meta"
@@ -219,7 +222,7 @@ let clenv_assign mv rhs clenv =
*)
let clenv_metas_in_type_of_meta evd mv =
- (mk_freelisted (meta_instance evd (meta_ftype evd mv))).freemetas
+ (mk_freelisted (meta_instance evd (map_fl EConstr.of_constr (meta_ftype evd mv)))).freemetas
let dependent_in_type_of_metas clenv mvs =
List.fold_right
@@ -262,35 +265,38 @@ let clenv_unify ?(flags=default_unify_flags ()) cv_pb t1 t2 clenv =
let clenv_unify_meta_types ?(flags=default_unify_flags ()) clenv =
{ clenv with evd = w_unify_meta_types ~flags:flags clenv.env clenv.evd }
-let clenv_unique_resolver ?(flags=default_unify_flags ()) clenv gl =
- let concl = Goal.V82.concl clenv.evd (sig_it gl) in
- if isMeta (fst (decompose_appvect (whd_nored clenv.evd clenv.templtyp.rebus))) then
+let clenv_unique_resolver_gen ?(flags=default_unify_flags ()) clenv concl =
+ if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (whd_nored clenv.evd clenv.templtyp.rebus))) then
clenv_unify CUMUL ~flags (clenv_type clenv) concl
(clenv_unify_meta_types ~flags clenv)
else
clenv_unify CUMUL ~flags
(meta_reducible_instance clenv.evd clenv.templtyp) concl clenv
+let old_clenv_unique_resolver ?flags clenv gl =
+ let concl = Goal.V82.concl clenv.evd (sig_it gl) in
+ clenv_unique_resolver_gen ?flags clenv concl
+
+let clenv_unique_resolver ?flags clenv gl =
+ let concl = Proofview.Goal.concl gl in
+ clenv_unique_resolver_gen ?flags clenv concl
+
let adjust_meta_source evd mv = function
| loc,Evar_kinds.VarInstance id ->
let rec match_name c l =
- match kind_of_term c, l with
- | Lambda (Name id,_,c), a::l when Constr.equal a (mkMeta mv) -> Some id
+ match EConstr.kind evd c, l with
+ | Lambda (Name id,_,c), a::l when EConstr.eq_constr evd a (mkMeta mv) -> Some id
| Lambda (_,_,c), a::l -> match_name c l
| _ -> None in
(* This is very ad hoc code so that an evar inherits the name of the binder
in situations like "ex_intro (fun x => P) ?ev p" *)
let f = function (mv',(Cltyp (_,t) | Clval (_,_,t))) ->
if Metaset.mem mv t.freemetas then
- let f,l = decompose_app t.rebus in
- match kind_of_term f with
+ let f,l = decompose_app evd (EConstr.of_constr t.rebus) in
+ match EConstr.kind evd f with
| Meta mv'' ->
(match meta_opt_fvalue evd mv'' with
- | Some (c,_) -> match_name c.rebus l
- | None -> None)
- | Evar ev ->
- (match existential_opt_value evd ev with
- | Some c -> match_name c l
+ | Some (c,_) -> match_name (EConstr.of_constr c.rebus) l
| None -> None)
| _ -> None
else None in
@@ -332,7 +338,7 @@ let clenv_pose_metas_as_evars clenv dep_mvs =
let ty = clenv_meta_type clenv mv in
(* Postpone the evar-ization if dependent on another meta *)
(* This assumes no cycle in the dependencies - is it correct ? *)
- if occur_meta ty then fold clenv (mvs@[mv])
+ if occur_meta clenv.evd ty then fold clenv (mvs@[mv])
else
let src = evar_source_of_meta mv clenv.evd in
let src = adjust_meta_source clenv.evd mv src in
@@ -404,7 +410,7 @@ type arg_bindings = constr explicit_bindings
* of cval, ctyp. *)
let clenv_independent clenv =
- let mvs = collect_metas (clenv_value clenv) in
+ let mvs = collect_metas clenv.evd (clenv_value clenv) in
let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in
let deps = Metaset.union (dependent_in_type_of_metas clenv mvs) ctyp_mvs in
List.filter (fun mv -> not (Metaset.mem mv deps)) mvs
@@ -482,7 +488,7 @@ let error_already_defined b =
(str "Position " ++ int n ++ str" already defined.")
let clenv_unify_binding_type clenv c t u =
- if isMeta (fst (decompose_appvect (whd_nored clenv.evd u))) then
+ if isMeta clenv.evd (fst (decompose_app_vect clenv.evd (whd_nored clenv.evd u))) then
(* Not enough information to know if some subtyping is needed *)
CoerceToType, clenv, c
else
@@ -501,6 +507,7 @@ let clenv_assign_binding clenv k c =
let k_typ = clenv_hnf_constr clenv (clenv_meta_type clenv k) in
let c_typ = nf_betaiota clenv.evd (clenv_get_type_of clenv c) in
let status,clenv',c = clenv_unify_binding_type clenv c c_typ k_typ in
+ let c = EConstr.Unsafe.to_constr c in
{ clenv' with evd = meta_assign k (c,(Conv,status)) clenv'.evd }
let clenv_match_args bl clenv =
@@ -513,7 +520,7 @@ let clenv_match_args bl clenv =
(fun clenv (loc,b,c) ->
let k = meta_of_binder clenv loc mvs b in
if meta_defined clenv.evd k then
- if Term.eq_constr (fst (meta_fvalue clenv.evd k)).rebus c then clenv
+ if EConstr.eq_constr clenv.evd (EConstr.of_constr (fst (meta_fvalue clenv.evd k)).rebus) c then clenv
else error_already_defined b
else
clenv_assign_binding clenv k c)
@@ -522,7 +529,7 @@ let clenv_match_args bl clenv =
exception NoSuchBinding
let clenv_constrain_last_binding c clenv =
- let all_mvs = collect_metas clenv.templval.rebus in
+ let all_mvs = collect_metas clenv.evd clenv.templval.rebus in
let k = try List.last all_mvs with Failure _ -> raise NoSuchBinding in
clenv_assign_binding clenv k c
@@ -557,8 +564,9 @@ let make_clenv_binding_gen hyps_only n env sigma (c,t) = function
let clause = mk_clenv_from_env env sigma n (c,t) in
clenv_constrain_dep_args hyps_only largs clause
| ExplicitBindings lbind ->
+ let t = rename_bound_vars_as_displayed sigma [] [] t in
let clause = mk_clenv_from_env env sigma n
- (c,rename_bound_vars_as_displayed [] [] t)
+ (c, t)
in clenv_match_args lbind clause
| NoBindings ->
mk_clenv_from_env env sigma n (c,t)
@@ -585,34 +593,36 @@ let pr_clenv clenv =
(** Evar version of mk_clenv *)
type hole = {
- hole_evar : constr;
- hole_type : types;
+ hole_evar : EConstr.constr;
+ hole_type : EConstr.types;
hole_deps : bool;
hole_name : Name.t;
}
type clause = {
cl_holes : hole list;
- cl_concl : types;
+ cl_concl : EConstr.types;
}
let make_evar_clause env sigma ?len t =
+ let open EConstr in
+ let open Vars in
let bound = match len with
| None -> -1
| Some n -> assert (0 <= n); n
in
(** FIXME: do the renaming online *)
- let t = rename_bound_vars_as_displayed [] [] t in
+ let t = rename_bound_vars_as_displayed sigma [] [] t in
let rec clrec (sigma, holes) n t =
if n = 0 then (sigma, holes, t)
- else match kind_of_term t with
+ else match EConstr.kind sigma t with
| Cast (t, _, _) -> clrec (sigma, holes) n t
| Prod (na, t1, t2) ->
let store = Typeclasses.set_resolvable Evd.Store.empty false in
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (ev, sigma, _) = new_evar ~store env sigma t1 in
let sigma = Sigma.to_evar_map sigma in
- let dep = dependent (mkRel 1) t2 in
+ let dep = not (noccurn sigma 1 t2) in
let hole = {
hole_evar = ev;
hole_type = t1;
@@ -667,24 +677,26 @@ let evar_of_binder holes = function
user_err (str "No such binder.")
let define_with_type sigma env ev c =
+ let open EConstr in
let t = Retyping.get_type_of env sigma ev in
let ty = Retyping.get_type_of env sigma c in
let j = Environ.make_judge c ty in
let (sigma, j) = Coercion.inh_conv_coerce_to true (Loc.ghost) env sigma j t in
- let (ev, _) = destEvar ev in
- let sigma = Evd.define ev j.Environ.uj_val sigma in
+ let (ev, _) = destEvar sigma ev in
+ let sigma = Evd.define ev (EConstr.Unsafe.to_constr j.Environ.uj_val) sigma in
sigma
let solve_evar_clause env sigma hyp_only clause = function
| NoBindings -> sigma
| ImplicitBindings largs ->
+ let open EConstr in
let fold holes h =
if h.hole_deps then
(** Some subsequent term uses the hole *)
- let (ev, _) = destEvar h.hole_evar in
- let is_dep hole = occur_evar ev hole.hole_type in
+ let (ev, _) = destEvar sigma h.hole_evar in
+ let is_dep hole = occur_evar sigma ev hole.hole_type in
let in_hyp = List.exists is_dep holes in
- let in_ccl = occur_evar ev clause.cl_concl in
+ let in_ccl = occur_evar sigma ev clause.cl_concl in
let dep = if hyp_only then in_hyp && not in_ccl else in_hyp || in_ccl in
let h = { h with hole_deps = dep } in
h :: holes
diff --git a/proofs/clenv.mli b/proofs/clenv.mli
index e9236b1da..4bcd50591 100644
--- a/proofs/clenv.mli
+++ b/proofs/clenv.mli
@@ -14,6 +14,7 @@ open Names
open Term
open Environ
open Evd
+open EConstr
open Unification
open Misctypes
@@ -28,8 +29,6 @@ type clausenv = {
templtyp : constr freelisted (** its type *)}
-val map_clenv : (constr -> constr) -> clausenv -> clausenv
-
(** subject of clenv (instantiated) *)
val clenv_value : clausenv -> constr
@@ -37,16 +36,16 @@ val clenv_value : clausenv -> constr
val clenv_type : clausenv -> types
(** substitute resolved metas *)
-val clenv_nf_meta : clausenv -> constr -> constr
+val clenv_nf_meta : clausenv -> EConstr.constr -> EConstr.constr
(** type of a meta in clenv context *)
val clenv_meta_type : clausenv -> metavariable -> types
-val mk_clenv_from : Goal.goal sigma -> constr * types -> clausenv
+val mk_clenv_from : ('a, 'r) Proofview.Goal.t -> EConstr.constr * EConstr.types -> clausenv
val mk_clenv_from_n :
- Goal.goal sigma -> int option -> constr * types -> clausenv
-val mk_clenv_type_of : Goal.goal sigma -> constr -> clausenv
-val mk_clenv_from_env : env -> evar_map -> int option -> constr * types -> clausenv
+ ('a, 'r) Proofview.Goal.t -> int option -> EConstr.constr * EConstr.types -> clausenv
+val mk_clenv_type_of : ('a, 'r) Proofview.Goal.t -> EConstr.constr -> clausenv
+val mk_clenv_from_env : env -> evar_map -> int option -> EConstr.constr * EConstr.types -> clausenv
(** Refresh the universes in a clenv *)
val refresh_undefined_univs : clausenv -> clausenv * Univ.universe_level_subst
@@ -63,9 +62,12 @@ val clenv_unify :
?flags:unify_flags -> conv_pb -> constr -> constr -> clausenv -> clausenv
(** unifies the concl of the goal with the type of the clenv *)
-val clenv_unique_resolver :
+val old_clenv_unique_resolver :
?flags:unify_flags -> clausenv -> Goal.goal sigma -> clausenv
+val clenv_unique_resolver :
+ ?flags:unify_flags -> clausenv -> ('a, 'r) Proofview.Goal.t -> clausenv
+
val clenv_dependent : clausenv -> metavariable list
val clenv_pose_metas_as_evars : clausenv -> metavariable list -> clausenv
@@ -92,18 +94,18 @@ val clenv_unify_meta_types : ?flags:unify_flags -> clausenv -> clausenv
the optional int tells how many prods of the lemma have to be used
use all of them if None *)
val make_clenv_binding_env_apply :
- env -> evar_map -> int option -> constr * constr -> constr bindings ->
+ env -> evar_map -> int option -> EConstr.constr * EConstr.constr -> constr bindings ->
clausenv
val make_clenv_binding_apply :
- env -> evar_map -> int option -> constr * constr -> constr bindings ->
+ env -> evar_map -> int option -> EConstr.constr * EConstr.constr -> constr bindings ->
clausenv
val make_clenv_binding_env :
- env -> evar_map -> constr * constr -> constr bindings -> clausenv
+ env -> evar_map -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv
val make_clenv_binding :
- env -> evar_map -> constr * constr -> constr bindings -> clausenv
+ env -> evar_map -> EConstr.constr * EConstr.constr -> constr bindings -> clausenv
(** if the clause is a product, add an extra meta for this product *)
exception NotExtensibleClause
@@ -134,9 +136,9 @@ val pr_clenv : clausenv -> Pp.std_ppcmds
*)
type hole = {
- hole_evar : constr;
+ hole_evar : EConstr.constr;
(** The hole itself. Guaranteed to be an evar. *)
- hole_type : types;
+ hole_type : EConstr.types;
(** Type of the hole in the current environment. *)
hole_deps : bool;
(** Whether the remainder of the clause was dependent in the hole. Note that
@@ -148,10 +150,10 @@ type hole = {
type clause = {
cl_holes : hole list;
- cl_concl : types;
+ cl_concl : EConstr.types;
}
-val make_evar_clause : env -> evar_map -> ?len:int -> types ->
+val make_evar_clause : env -> evar_map -> ?len:int -> EConstr.types ->
(evar_map * clause)
(** An evar version of {!make_clenv_binding}. Given a type [t],
[evar_environments env sigma ~len t bl] tries to eliminate at most [len]
@@ -159,7 +161,7 @@ val make_evar_clause : env -> evar_map -> ?len:int -> types ->
type together with the list of holes generated. Assumes that [t] is
well-typed in the environment. *)
-val solve_evar_clause : env -> evar_map -> bool -> clause -> constr bindings ->
+val solve_evar_clause : env -> evar_map -> bool -> clause -> EConstr.constr bindings ->
evar_map
(** [solve_evar_clause env sigma hyps cl bl] tries to solve the holes contained
in [cl] according to the [bl] argument. Assumes that [bl] are well-typed in
diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml
index 98b5bc8b0..0722ea047 100644
--- a/proofs/clenvtac.ml
+++ b/proofs/clenvtac.ml
@@ -11,6 +11,7 @@ open Names
open Term
open Termops
open Evd
+open EConstr
open Refiner
open Logic
open Reduction
@@ -26,19 +27,19 @@ open Proofview.Notations
let clenv_cast_meta clenv =
let rec crec u =
- match kind_of_term u with
+ match EConstr.kind clenv.evd u with
| App _ | Case _ -> crec_hd u
- | Cast (c,_,_) when isMeta c -> u
+ | Cast (c,_,_) when isMeta clenv.evd c -> u
| Proj (p, c) -> mkProj (p, crec_hd c)
- | _ -> map_constr crec u
+ | _ -> EConstr.map clenv.evd crec u
and crec_hd u =
- match kind_of_term (strip_outer_cast u) with
+ match EConstr.kind clenv.evd (strip_outer_cast clenv.evd u) with
| Meta mv ->
(try
let b = Typing.meta_type clenv.evd mv in
- assert (not (occur_meta b));
- if occur_meta b then u
+ assert (not (occur_meta clenv.evd b));
+ if occur_meta clenv.evd b then u
else mkCast (mkMeta mv, DEFAULTcast, b)
with Not_found -> u)
| App(f,args) -> mkApp (crec_hd f, Array.map crec args)
@@ -104,8 +105,8 @@ let dft = default_unify_flags
let res_pf ?(with_evars=false) ?(with_classes=true) ?(flags=dft ()) clenv =
Proofview.Goal.enter { enter = begin fun gl ->
- let clenv gl = clenv_unique_resolver ~flags clenv gl in
- clenv_refine with_evars ~with_classes (Tacmach.New.of_old clenv (Proofview.Goal.assume gl))
+ let clenv = clenv_unique_resolver ~flags clenv gl in
+ clenv_refine with_evars ~with_classes clenv
end }
(* [unifyTerms] et [unify] ne semble pas gérer les Meta, en
diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli
index 8a096b645..5b7164705 100644
--- a/proofs/clenvtac.mli
+++ b/proofs/clenvtac.mli
@@ -10,6 +10,7 @@
open Term
open Clenv
+open EConstr
open Unification
open Misctypes
diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml
index 509efbc5b..8367c09b8 100644
--- a/proofs/evar_refiner.ml
+++ b/proofs/evar_refiner.ml
@@ -18,21 +18,23 @@ open Pp
(* Instantiation of existential variables *)
(******************************************)
-let depends_on_evar evk _ (pbty,_,t1,t2) =
- try Evar.equal (head_evar t1) evk
+let depends_on_evar sigma evk _ (pbty,_,t1,t2) =
+ let t1 = EConstr.of_constr t1 in
+ let t2 = EConstr.of_constr t2 in
+ try Evar.equal (head_evar sigma t1) evk
with NoHeadEvar ->
- try Evar.equal (head_evar t2) evk
+ try Evar.equal (head_evar sigma t2) evk
with NoHeadEvar -> false
let define_and_solve_constraints evk c env evd =
- if Termops.occur_evar evk c then
+ if Termops.occur_evar evd evk c then
Pretype_errors.error_occur_check env evd evk c;
- let evd = define evk c evd in
- let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evk) in
+ let evd = define evk (EConstr.Unsafe.to_constr c) evd in
+ let (evd,pbs) = extract_changed_conv_pbs evd (depends_on_evar evd evk) in
match
List.fold_left
(fun p (pbty,env,t1,t2) -> match p with
- | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty t1 t2
+ | Success evd -> Evarconv.evar_conv_x full_transparent_state env evd pbty (EConstr.of_constr t1) (EConstr.of_constr t2)
| UnifFailure _ as x -> x) (Success evd)
pbs
with
@@ -51,11 +53,11 @@ let w_refine (evk,evi) (ltac_var,rawc) sigma =
Pretyping.fail_evar = false;
Pretyping.expand_evars = true } in
try Pretyping.understand_ltac flags
- env sigma ltac_var (Pretyping.OfType evi.evar_concl) rawc
+ env sigma ltac_var (Pretyping.OfType (EConstr.of_constr evi.evar_concl)) rawc
with e when CErrors.noncritical e ->
let loc = Glob_ops.loc_of_glob_constr rawc in
user_err ~loc
(str "Instance is not well-typed in the environment of " ++
- pr_existential_key sigma evk ++ str ".")
+ Termops.pr_existential_key sigma evk ++ str ".")
in
define_and_solve_constraints evk typed_c env (evars_reset_evd sigma' sigma)
diff --git a/proofs/goal.ml b/proofs/goal.ml
index a141708c2..9046f4534 100644
--- a/proofs/goal.ml
+++ b/proofs/goal.ml
@@ -49,7 +49,7 @@ module V82 = struct
(* Access to ".evar_concl" *)
let concl evars gl =
let evi = Evd.find evars gl in
- evi.Evd.evar_concl
+ EConstr.of_constr evi.Evd.evar_concl
(* Access to ".evar_extra" *)
let extra evars gl =
@@ -62,6 +62,7 @@ module V82 = struct
be shelved. It must not appear as a future_goal, so the future
goals are restored to their initial value after the evar is
created. *)
+ let concl = EConstr.Unsafe.to_constr concl in
let prev_future_goals = Evd.future_goals evars in
let prev_principal_goal = Evd.principal_future_goal evars in
let evi = { Evd.evar_hyps = hyps;
@@ -78,8 +79,8 @@ module V82 = struct
let evars = Sigma.to_evar_map evars in
let evars = Evd.restore_future_goals evars prev_future_goals prev_principal_goal in
let ctxt = Environ.named_context_of_val hyps in
- let inst = Array.map_of_list (NamedDecl.get_id %> mkVar) ctxt in
- let ev = Term.mkEvar (evk,inst) in
+ let inst = Array.map_of_list (NamedDecl.get_id %> EConstr.mkVar) ctxt in
+ let ev = EConstr.mkEvar (evk,inst) in
(evk, ev, evars)
(* Instantiates a goal with an open term *)
@@ -89,7 +90,7 @@ module V82 = struct
if not (Evarutil.occur_evar_upto sigma evk c) then ()
else Pretype_errors.error_occur_check Environ.empty_env sigma evk c
in
- Evd.define evk c sigma
+ Evd.define evk (EConstr.Unsafe.to_constr c) sigma
(* Instantiates a goal with an open term, using name of goal for evk' *)
let partial_solution_to sigma evk evk' c =
@@ -145,6 +146,7 @@ module V82 = struct
(* Goal represented as a type, doesn't take into account section variables *)
let abstract_type sigma gl =
+ let open EConstr in
let (gl,sigma) = nf_evar sigma gl in
let env = env sigma gl in
let genv = Global.env () in
@@ -153,6 +155,7 @@ module V82 = struct
with Not_found -> true in
Environ.fold_named_context_reverse (fun t decl ->
if is_proof_var decl then
+ let decl = Termops.map_named_decl EConstr.of_constr decl in
mkNamedProd_or_LetIn decl t
else
t
diff --git a/proofs/goal.mli b/proofs/goal.mli
index 6a79c1f45..a2fa34c05 100644
--- a/proofs/goal.mli
+++ b/proofs/goal.mli
@@ -38,7 +38,7 @@ module V82 : sig
val nf_hyps : Evd.evar_map -> goal -> Environ.named_context_val
(* Access to ".evar_concl" *)
- val concl : Evd.evar_map -> goal -> Term.constr
+ val concl : Evd.evar_map -> goal -> EConstr.constr
(* Access to ".evar_extra" *)
val extra : Evd.evar_map -> goal -> Evd.Store.t
@@ -48,16 +48,16 @@ module V82 : sig
the evar corresponding to the goal, and an updated evar_map. *)
val mk_goal : Evd.evar_map ->
Environ.named_context_val ->
- Term.constr ->
+ EConstr.constr ->
Evd.Store.t ->
- goal * Term.constr * Evd.evar_map
+ goal * EConstr.constr * Evd.evar_map
(* Instantiates a goal with an open term *)
- val partial_solution : Evd.evar_map -> goal -> Term.constr -> Evd.evar_map
+ val partial_solution : Evd.evar_map -> goal -> EConstr.constr -> Evd.evar_map
(* Instantiates a goal with an open term, reusing name of goal for
second goal *)
- val partial_solution_to : Evd.evar_map -> goal -> goal -> Term.constr -> Evd.evar_map
+ val partial_solution_to : Evd.evar_map -> goal -> goal -> EConstr.constr -> Evd.evar_map
(* Principal part of the weak-progress tactical *)
val weak_progress : goal list Evd.sigma -> goal Evd.sigma -> bool
@@ -75,6 +75,6 @@ module V82 : sig
val nf_evar : Evd.evar_map -> goal -> goal * Evd.evar_map
(* Goal represented as a type, doesn't take into account section variables *)
- val abstract_type : Evd.evar_map -> goal -> Term.types
+ val abstract_type : Evd.evar_map -> goal -> EConstr.types
end
diff --git a/proofs/logic.ml b/proofs/logic.ml
index 44c629484..e6024785d 100644
--- a/proofs/logic.ml
+++ b/proofs/logic.ml
@@ -34,7 +34,7 @@ type refiner_error =
| CannotApply of constr * constr
| NotWellTyped of constr
| NonLinearProof of constr
- | MetaInType of constr
+ | MetaInType of EConstr.constr
(* Errors raised by the tactics *)
| IntroNeedsProduct
@@ -86,7 +86,7 @@ let apply_to_hyp check sign id f =
else sign
let check_typability env sigma c =
- if !check then let _ = unsafe_type_of env sigma c in ()
+ if !check then let _ = unsafe_type_of env sigma (EConstr.of_constr c) in ()
(************************************************************************)
(************************************************************************)
@@ -138,12 +138,12 @@ let find_q x (m,q) =
else find (Id.Set.union l accs) (i::acc) itl in
find Id.Set.empty [] q
-let occur_vars_in_decl env hyps d =
+let occur_vars_in_decl env sigma hyps d =
if Id.Set.is_empty hyps then false else
- let ohyps = global_vars_set_of_decl env d in
+ let ohyps = global_vars_set_of_decl env sigma d in
Id.Set.exists (fun h -> Id.Set.mem h ohyps) hyps
-let reorder_context env sign ord =
+let reorder_context env sigma sign ord =
let ords = List.fold_right Id.Set.add ord Id.Set.empty in
if not (Int.equal (List.length ord) (Id.Set.cardinal ords)) then
error "Order list has duplicates";
@@ -152,13 +152,13 @@ let reorder_context env sign ord =
| [] -> List.rev ctxt_tail @ ctxt_head
| top::ord' when mem_q top moved_hyps ->
let ((d,h),mh) = find_q top moved_hyps in
- if occur_vars_in_decl env h d then
+ if occur_vars_in_decl env sigma h d then
user_err ~hdr:"reorder_context"
(str "Cannot move declaration " ++ pr_id top ++ spc() ++
str "before " ++
pr_sequence pr_id
(Id.Set.elements (Id.Set.inter h
- (global_vars_set_of_decl env d))));
+ (global_vars_set_of_decl env sigma d))));
step ord' expected ctxt_head mh (d::ctxt_tail)
| _ ->
(match ctxt_head with
@@ -173,16 +173,18 @@ let reorder_context env sign ord =
ctxt (push_val x moved_hyps) (d::ctxt_tail)) in
step ord ords sign mt_q []
-let reorder_val_context env sign ord =
- val_of_named_context (reorder_context env (named_context_of_val sign) ord)
+let reorder_val_context env sigma sign ord =
+ let open EConstr in
+ val_of_named_context (reorder_context env sigma (named_context_of_val sign) ord)
-let check_decl_position env sign d =
+let check_decl_position env sigma sign d =
+ let open EConstr in
let x = NamedDecl.get_id d in
- let needed = global_vars_set_of_decl env d in
- let deps = dependency_closure env (named_context_of_val sign) needed in
+ let needed = global_vars_set_of_decl env sigma d in
+ let deps = dependency_closure env sigma (named_context_of_val sign) needed in
if Id.List.mem x deps then
user_err ~hdr:"Logic.check_decl_position"
(str "Cannot create self-referring hypothesis " ++ pr_id x);
@@ -233,12 +235,12 @@ let hyp_of_move_location = function
| MoveBefore id -> id
| _ -> assert false
-let move_hyp toleft (left,declfrom,right) hto =
+let move_hyp sigma toleft (left,declfrom,right) hto =
let env = Global.env() in
let test_dep d d2 =
if toleft
- then occur_var_in_decl env (NamedDecl.get_id d2) d
- else occur_var_in_decl env (NamedDecl.get_id d) d2
+ then occur_var_in_decl env sigma (NamedDecl.get_id d2) d
+ else occur_var_in_decl env sigma (NamedDecl.get_id d) d2
in
let rec moverec first middle = function
| [] ->
@@ -266,6 +268,7 @@ let move_hyp toleft (left,declfrom,right) hto =
else
moverec first' middle' right
in
+ let open EConstr in
if toleft then
let right =
List.fold_right push_named_context_val right empty_named_context_val in
@@ -278,10 +281,11 @@ let move_hyp toleft (left,declfrom,right) hto =
List.fold_left (fun sign d -> push_named_context_val d sign)
right left
-let move_hyp_in_named_context hfrom hto sign =
+let move_hyp_in_named_context sigma hfrom hto sign =
+ let open EConstr in
let (left,right,declfrom,toleft) =
split_sign hfrom hto (named_context_of_val sign) in
- move_hyp toleft (left,declfrom,right) hto
+ move_hyp sigma toleft (left,declfrom,right) hto
(**********************************************************************)
@@ -314,23 +318,26 @@ let check_meta_variables c =
let check_conv_leq_goal env sigma arg ty conclty =
if !check then
- let evm, b = Reductionops.infer_conv env sigma ty conclty in
+ let evm, b = Reductionops.infer_conv env sigma (EConstr.of_constr ty) (EConstr.of_constr conclty) in
if b then evm
else raise (RefinerError (BadType (arg,ty,conclty)))
else sigma
-exception Stop of constr list
-let meta_free_prefix a =
+exception Stop of EConstr.t list
+let meta_free_prefix sigma a =
try
+ let a = Array.map EConstr.of_constr a in
let _ = Array.fold_left (fun acc a ->
- if occur_meta a then raise (Stop acc)
+ if occur_meta sigma a then raise (Stop acc)
else a :: acc) [] a
in a
with Stop acc -> Array.rev_of_list acc
let goal_type_of env sigma c =
- if !check then unsafe_type_of env sigma c
- else Retyping.get_type_of env sigma c
+ if !check then
+ let (sigma,t) = type_of env sigma (EConstr.of_constr c) in
+ (sigma, EConstr.Unsafe.to_constr t)
+ else (sigma, EConstr.Unsafe.to_constr (Retyping.get_type_of env sigma (EConstr.of_constr c)))
let rec mk_refgoals sigma goal goalacc conclty trm =
let env = Goal.V82.env sigma goal in
@@ -338,17 +345,20 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
let mk_goal hyps concl =
Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal)
in
- if (not !check) && not (occur_meta trm) then
- let t'ty = Retyping.get_type_of env sigma trm in
+ if (not !check) && not (occur_meta sigma (EConstr.of_constr trm)) then
+ let t'ty = Retyping.get_type_of env sigma (EConstr.of_constr trm) in
+ let t'ty = EConstr.Unsafe.to_constr t'ty in
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
| Meta _ ->
- let conclty = nf_betaiota sigma conclty in
- if !check && occur_meta conclty then
+ let conclty = nf_betaiota sigma (EConstr.of_constr conclty) in
+ if !check && occur_meta sigma conclty then
raise (RefinerError (MetaInType conclty));
let (gl,ev,sigma) = mk_goal hyps conclty in
+ let ev = EConstr.Unsafe.to_constr ev in
+ let conclty = EConstr.Unsafe.to_constr conclty in
gl::goalacc, conclty, sigma, ev
| Cast (t,k, ty) ->
@@ -367,13 +377,14 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
| App (f,l) ->
let (acc',hdty,sigma,applicand) =
- if is_template_polymorphic env f then
+ if is_template_polymorphic env sigma (EConstr.of_constr f) then
let ty =
(* Template sort-polymorphism of definition and inductive types *)
- let firstmeta = Array.findi (fun i x -> occur_meta x) l in
+ let firstmeta = Array.findi (fun i x -> occur_meta sigma (EConstr.of_constr x)) l in
let args, _ = Option.cata (fun i -> CArray.chop i l) (l, [||]) firstmeta in
- type_of_global_reference_knowing_parameters env sigma f args
+ type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) (Array.map EConstr.of_constr args)
in
+ let ty = EConstr.Unsafe.to_constr ty in
goalacc, ty, sigma, f
else
mk_hdgoals sigma goal goalacc f
@@ -386,7 +397,8 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
| Proj (p,c) ->
let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in
let c = mkProj (p, c') in
- let ty = get_type_of env sigma c in
+ let ty = get_type_of env sigma (EConstr.of_constr c) in
+ let ty = EConstr.Unsafe.to_constr ty in
(acc',ty,sigma,c)
| Case (ci,p,c,lf) ->
@@ -406,9 +418,9 @@ let rec mk_refgoals sigma goal goalacc conclty trm =
(acc'',conclty',sigma, ans)
| _ ->
- if occur_meta trm then
+ if occur_meta sigma (EConstr.of_constr trm) then
anomaly (Pp.str "refiner called with a meta in non app/case subterm");
- let t'ty = goal_type_of env sigma trm in
+ let (sigma, t'ty) = goal_type_of env sigma trm in
let sigma = check_conv_leq_goal env sigma trm t'ty conclty in
(goalacc,t'ty,sigma, trm)
@@ -423,7 +435,8 @@ and mk_hdgoals sigma goal goalacc trm =
match kind_of_term trm with
| Cast (c,_, ty) when isMeta c ->
check_typability env sigma ty;
- let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma ty) in
+ let (gl,ev,sigma) = mk_goal hyps (nf_betaiota sigma (EConstr.of_constr ty)) in
+ let ev = EConstr.Unsafe.to_constr ev in
gl::goalacc,ty,sigma,ev
| Cast (t,_, ty) ->
@@ -432,10 +445,10 @@ and mk_hdgoals sigma goal goalacc trm =
| App (f,l) ->
let (acc',hdty,sigma,applicand) =
- if is_template_polymorphic env f
+ if is_template_polymorphic env sigma (EConstr.of_constr f)
then
- let l' = meta_free_prefix l in
- (goalacc,type_of_global_reference_knowing_parameters env sigma f l',sigma,f)
+ let l' = meta_free_prefix sigma l in
+ (goalacc,EConstr.Unsafe.to_constr (type_of_global_reference_knowing_parameters env sigma (EConstr.of_constr f) l'),sigma,f)
else mk_hdgoals sigma goal goalacc f
in
let ((acc'',conclty',sigma), args) = mk_arggoals sigma goal acc' hdty l in
@@ -460,17 +473,20 @@ and mk_hdgoals sigma goal goalacc trm =
| Proj (p,c) ->
let (acc',cty,sigma,c') = mk_hdgoals sigma goal goalacc c in
let c = mkProj (p, c') in
- let ty = get_type_of env sigma c in
+ let ty = get_type_of env sigma (EConstr.of_constr c) in
+ let ty = EConstr.Unsafe.to_constr ty in
(acc',ty,sigma,c)
| _ ->
- if !check && occur_meta trm then
+ if !check && occur_meta sigma (EConstr.of_constr trm) then
anomaly (Pp.str "refine called with a dependent meta");
- goalacc, goal_type_of env sigma trm, sigma, trm
+ let (sigma, ty) = goal_type_of env sigma trm in
+ goalacc, ty, sigma, trm
and mk_arggoals sigma goal goalacc funty allargs =
let foldmap (goalacc, funty, sigma) harg =
- let t = whd_all (Goal.V82.env sigma goal) sigma funty in
+ 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
| LetIn (_, c1, _, b) -> collapse (subst1 c1 b)
| _ -> t
@@ -487,11 +503,13 @@ and mk_arggoals sigma goal goalacc funty allargs =
and mk_casegoals sigma goal goalacc p c =
let env = Goal.V82.env sigma goal in
let (acc',ct,sigma,c') = mk_hdgoals sigma goal goalacc c in
+ let ct = EConstr.of_constr ct in
let (acc'',pt,sigma,p') = mk_hdgoals sigma goal acc' p in
- let indspec =
+ let ((ind, u), spec) =
try Tacred.find_hnf_rectype env sigma ct
with Not_found -> anomaly (Pp.str "mk_casegoals") in
- let (lbrty,conclty) = type_case_branches_with_names env indspec p c in
+ let indspec = ((ind, EConstr.EInstance.kind sigma u), spec) in
+ let (lbrty,conclty) = type_case_branches_with_names env sigma indspec p c in
(acc'',lbrty,conclty,sigma,p',c')
@@ -503,17 +521,17 @@ let convert_hyp check sign sigma d =
let sign' =
apply_to_hyp check sign id
(fun _ d' _ ->
- let c = NamedDecl.get_value d' in
+ let c = Option.map EConstr.of_constr (NamedDecl.get_value d') in
let env = Global.env_of_context sign in
- if check && not (is_conv env sigma (NamedDecl.get_type d) (NamedDecl.get_type d')) then
+ if check && not (is_conv env sigma (NamedDecl.get_type d) (EConstr.of_constr (NamedDecl.get_type d'))) then
user_err ~hdr:"Logic.convert_hyp"
(str "Incorrect change of the type of " ++ pr_id id ++ str ".");
if check && not (Option.equal (is_conv env sigma) b c) then
user_err ~hdr:"Logic.convert_hyp"
(str "Incorrect change of the body of "++ pr_id id ++ str ".");
- if check then reorder := check_decl_position env sign d;
- d) in
- reorder_val_context env sign' !reorder
+ if check then reorder := check_decl_position env sigma sign d;
+ map_named_decl EConstr.Unsafe.to_constr d) in
+ reorder_val_context env sigma sign' !reorder
@@ -528,16 +546,18 @@ let prim_refiner r sigma goal =
let mk_goal hyps concl =
Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal)
in
+ let open EConstr in
match r with
(* Logical rules *)
| Cut (b,replace,id,t) ->
(* if !check && not (Retyping.get_sort_of env sigma t) then*)
+ let t = EConstr.of_constr t in
let (sg1,ev1,sigma) = mk_goal sign (nf_betaiota sigma t) in
let sign,t,cl,sigma =
if replace then
let nexthyp = get_hyp_after id (named_context_of_val sign) in
let sign,t,cl,sigma = clear_hyps2 env sigma (Id.Set.singleton id) sign t cl in
- move_hyp false ([], LocalAssum (id,t),named_context_of_val sign)
+ move_hyp sigma false ([], LocalAssum (id,t),named_context_of_val sign)
nexthyp,
t,cl,sigma
else
@@ -547,13 +567,14 @@ let prim_refiner r sigma goal =
push_named_context_val (LocalAssum (id,t)) sign,t,cl,sigma) in
let (sg2,ev2,sigma) =
Goal.V82.mk_goal sigma sign cl (Goal.V82.extra sigma goal) in
- let oterm = Term.mkNamedLetIn id ev1 t ev2 in
+ let oterm = mkLetIn (Name id, ev1, t, EConstr.Vars.subst_var id ev2) in
let sigma = Goal.V82.partial_solution_to sigma goal sg2 oterm in
if b then ([sg1;sg2],sigma) else ([sg2;sg1],sigma)
| Refine c ->
+ let cl = EConstr.Unsafe.to_constr cl in
check_meta_variables c;
let (sgl,cl',sigma,oterm) = mk_refgoals sigma goal [] cl c in
let sgl = List.rev sgl in
- let sigma = Goal.V82.partial_solution sigma goal oterm in
+ let sigma = Goal.V82.partial_solution sigma goal (EConstr.of_constr oterm) in
(sgl, sigma)
diff --git a/proofs/logic.mli b/proofs/logic.mli
index 0dba9ef1e..f98248e4d 100644
--- a/proofs/logic.mli
+++ b/proofs/logic.mli
@@ -43,7 +43,7 @@ type refiner_error =
| CannotApply of constr * constr
| NotWellTyped of constr
| NonLinearProof of constr
- | MetaInType of constr
+ | MetaInType of EConstr.constr
(*i Errors raised by the tactics i*)
| IntroNeedsProduct
@@ -55,7 +55,7 @@ exception RefinerError of refiner_error
val catchable_exception : exn -> bool
val convert_hyp : bool -> Environ.named_context_val -> evar_map ->
- Context.Named.Declaration.t -> Environ.named_context_val
+ EConstr.named_declaration -> Environ.named_context_val
-val move_hyp_in_named_context : Id.t -> Id.t Misctypes.move_location ->
+val move_hyp_in_named_context : Evd.evar_map -> Id.t -> Id.t Misctypes.move_location ->
Environ.named_context_val -> Environ.named_context_val
diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml
index 9995a9394..7e8ed31d1 100644
--- a/proofs/pfedit.ml
+++ b/proofs/pfedit.ml
@@ -210,6 +210,7 @@ let refine_by_tactic env sigma ty tac =
| _ -> assert false
in
let ans = Reductionops.nf_evar sigma ans in
+ let ans = EConstr.Unsafe.to_constr ans in
(** [neff] contains the freshly generated side-effects *)
let neff = Evd.eval_side_effects sigma in
(** Reset the old side-effects *)
@@ -238,16 +239,17 @@ let apply_implicit_tactic tac = (); fun env sigma evk ->
match snd (evar_source evk sigma) with
| (Evar_kinds.ImplicitArg _ | Evar_kinds.QuestionMark _)
when
- Context.Named.equal (Environ.named_context_of_val evi.evar_hyps)
+ Context.Named.equal Constr.equal (Environ.named_context_of_val evi.evar_hyps)
(Environ.named_context env) ->
let tac = Proofview.tclTHEN tac (Proofview.tclEXTEND [] (Proofview.tclZERO (CErrors.UserError (None,Pp.str"Proof is not complete."))) []) in
(try
let c = Evarutil.nf_evars_universes sigma evi.evar_concl in
+ let c = EConstr.of_constr c in
if Evarutil.has_undefined_evars sigma c then raise Exit;
let (ans, _, ctx) =
build_by_tactic env (Evd.evar_universe_context sigma) c tac in
let sigma = Evd.set_universe_context sigma ctx in
- sigma, ans
+ sigma, EConstr.of_constr ans
with e when Logic.catchable_exception e -> raise Exit)
| _ -> raise Exit
diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli
index aad719db4..f9fb0b76d 100644
--- a/proofs/pfedit.mli
+++ b/proofs/pfedit.mli
@@ -60,7 +60,7 @@ type lemma_possible_guards = Proof_global.lemma_possible_guards
type universe_binders = Id.t Loc.located list
val start_proof :
- Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> constr ->
+ Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map -> named_context_val -> EConstr.constr ->
?init_tac:unit Proofview.tactic ->
Proof_global.proof_terminator -> unit
@@ -106,7 +106,7 @@ val get_current_context : unit -> Evd.evar_map * env
(** [current_proof_statement] *)
val current_proof_statement :
- unit -> Id.t * goal_kind * types
+ unit -> Id.t * goal_kind * EConstr.types
(** {6 ... } *)
(** [get_current_proof_name ()] return the name of the current focused
@@ -166,15 +166,15 @@ val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> unit
val build_constant_by_tactic :
Id.t -> Evd.evar_universe_context -> named_context_val -> ?goal_kind:goal_kind ->
- types -> unit Proofview.tactic ->
+ EConstr.types -> unit Proofview.tactic ->
Safe_typing.private_constants Entries.definition_entry * bool *
Evd.evar_universe_context
val build_by_tactic : ?side_eff:bool -> env -> Evd.evar_universe_context -> ?poly:polymorphic ->
- types -> unit Proofview.tactic ->
+ EConstr.types -> unit Proofview.tactic ->
constr * bool * Evd.evar_universe_context
-val refine_by_tactic : env -> Evd.evar_map -> types -> unit Proofview.tactic ->
+val refine_by_tactic : env -> Evd.evar_map -> EConstr.types -> unit Proofview.tactic ->
constr * Evd.evar_map
(** A variant of the above function that handles open terms as well.
Caveat: all effects are purged in the returned term at the end, but other
@@ -190,4 +190,4 @@ val declare_implicit_tactic : unit Proofview.tactic -> unit
val clear_implicit_tactic : unit -> unit
(* Raise Exit if cannot solve *)
-val solve_by_implicit_tactic : unit -> (env -> Evd.evar_map -> Evd.evar -> Evd.evar_map * constr) option
+val solve_by_implicit_tactic : unit -> (env -> Evd.evar_map -> Evd.evar -> Evd.evar_map * EConstr.constr) option
diff --git a/proofs/proof.mli b/proofs/proof.mli
index 8dc165e72..9d38e16ad 100644
--- a/proofs/proof.mli
+++ b/proofs/proof.mli
@@ -66,9 +66,9 @@ val map_structured_proof : proof -> (Evd.evar_map -> Goal.goal -> 'a) -> ('a pre
(*** General proof functions ***)
-val start : Evd.evar_map -> (Environ.env * Term.types) list -> proof
+val start : Evd.evar_map -> (Environ.env * EConstr.types) list -> proof
val dependent_start : Proofview.telescope -> proof
-val initial_goals : proof -> (Term.constr * Term.types) list
+val initial_goals : proof -> (EConstr.constr * EConstr.types) list
val initial_euctx : proof -> Evd.evar_universe_context
(* Returns [true] if the considered proof is completed, that is if no goal remain
@@ -79,7 +79,7 @@ val is_done : proof -> bool
val is_complete : proof -> bool
(* Returns the list of partial proofs to initial goals. *)
-val partial_proof : proof -> Term.constr list
+val partial_proof : proof -> EConstr.constr list
val compact : proof -> proof
diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml
index ca7330fdb..e4cba6f07 100644
--- a/proofs/proof_global.ml
+++ b/proofs/proof_global.ml
@@ -378,6 +378,7 @@ let close_proof ~keep_body_ucst_separate ?feedback_id ~now
in
let entries =
Future.map2 (fun p (_, t) ->
+ let t = EConstr.Unsafe.to_constr t in
let univstyp, body = make_body t p in
let univs, typ = Future.force univstyp in
{ Entries.
@@ -407,7 +408,7 @@ let return_proof ?(allow_partial=false) () =
(** ppedrot: FIXME, this is surely wrong. There is no reason to duplicate
side-effects... This may explain why one need to uniquize side-effects
thereafter... *)
- let proofs = List.map (fun c -> c, eff) proofs in
+ let proofs = List.map (fun c -> EConstr.Unsafe.to_constr c, eff) proofs in
proofs, Evd.evar_universe_context evd
end else
let initial_goals = Proof.initial_goals proof in
@@ -431,7 +432,7 @@ let return_proof ?(allow_partial=false) () =
side-effects... This may explain why one need to uniquize side-effects
thereafter... *)
let proofs =
- List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd c, eff)) initial_goals in
+ List.map (fun (c, _) -> (Evarutil.nf_evars_universes evd (EConstr.Unsafe.to_constr c), eff)) initial_goals in
proofs, Evd.evar_universe_context evd
let close_future_proof ~feedback_id proof =
diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli
index 97a21cf22..3b2cc6b23 100644
--- a/proofs/proof_global.mli
+++ b/proofs/proof_global.mli
@@ -86,7 +86,7 @@ val apply_terminator : proof_terminator -> proof_ending -> unit
end of the proof to close the proof. *)
val start_proof :
Evd.evar_map -> Names.Id.t -> ?pl:universe_binders ->
- Decl_kinds.goal_kind -> (Environ.env * Term.types) list ->
+ Decl_kinds.goal_kind -> (Environ.env * EConstr.types) list ->
proof_terminator -> unit
(** Like [start_proof] except that there may be dependencies between
@@ -201,7 +201,7 @@ end
val get_default_goal_selector : unit -> Vernacexpr.goal_selector
module V82 : sig
- val get_current_initial_conclusions : unit -> Names.Id.t *(Term.types list *
+ val get_current_initial_conclusions : unit -> Names.Id.t *(EConstr.types list *
Decl_kinds.goal_kind)
end
diff --git a/proofs/redexpr.ml b/proofs/redexpr.ml
index 34443b93d..0fe5c73f1 100644
--- a/proofs/redexpr.ml
+++ b/proofs/redexpr.ml
@@ -11,6 +11,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Declarations
open Globnames
open Genredexpr
@@ -25,9 +26,7 @@ open Misctypes
(* call by value normalisation function using the virtual machine *)
let cbv_vm env sigma c =
let ctyp = Retyping.get_type_of env sigma c in
- if Termops.occur_meta_or_existential c then
- error "vm_compute does not support existential variables.";
- Vnorm.cbv_vm env c ctyp
+ Vnorm.cbv_vm env sigma c ctyp
let warn_native_compute_disabled =
CWarnings.create ~name:"native-compute-disabled" ~category:"native-compiler"
@@ -45,7 +44,8 @@ let cbv_native env sigma c =
let whd_cbn flags env sigma t =
let (state,_) =
(whd_state_gen true true flags env sigma (t,Reductionops.Stack.empty))
- in Reductionops.Stack.zip ~refold:true state
+ in
+ Reductionops.Stack.zip ~refold:true sigma state
let strong_cbn flags =
strong (whd_cbn flags)
@@ -254,9 +254,12 @@ let reduction_of_red_expr env =
in
reduction_of_red_expr
+let subst_mps subst c =
+ EConstr.of_constr (Mod_subst.subst_mps subst (EConstr.Unsafe.to_constr c))
+
let subst_red_expr subs =
Miscops.map_red_expr_gen
- (Mod_subst.subst_mps subs)
+ (subst_mps subs)
(Mod_subst.subst_evaluable_reference subs)
(Patternops.subst_pattern subs)
diff --git a/proofs/redexpr.mli b/proofs/redexpr.mli
index d4c2c4a6c..45e461066 100644
--- a/proofs/redexpr.mli
+++ b/proofs/redexpr.mli
@@ -10,6 +10,7 @@
open Names
open Term
+open EConstr
open Pattern
open Genredexpr
open Reductionops
diff --git a/proofs/refine.ml b/proofs/refine.ml
index c238f731d..1ee6e0ca5 100644
--- a/proofs/refine.ml
+++ b/proofs/refine.ml
@@ -28,12 +28,12 @@ let typecheck_evar ev env sigma =
let info = Evd.find sigma ev in
(** Typecheck the hypotheses. *)
let type_hyp (sigma, env) decl =
- let t = NamedDecl.get_type decl in
+ let t = EConstr.of_constr (NamedDecl.get_type decl) in
let evdref = ref sigma in
let _ = Typing.e_sort_of env evdref t in
let () = match decl with
| LocalAssum _ -> ()
- | LocalDef (_,body,_) -> Typing.e_check env evdref body t
+ | LocalDef (_,body,_) -> Typing.e_check env evdref (EConstr.of_constr body) t
in
(!evdref, Environ.push_named decl env)
in
@@ -42,7 +42,7 @@ let typecheck_evar ev env sigma =
let (sigma, env) = List.fold_left type_hyp (sigma, env) changed in
(** Typecheck the conclusion *)
let evdref = ref sigma in
- let _ = Typing.e_sort_of env evdref (Evd.evar_concl info) in
+ let _ = Typing.e_sort_of env evdref (EConstr.of_constr (Evd.evar_concl info)) in
!evdref
let typecheck_proof c concl env sigma =
@@ -101,6 +101,7 @@ let make_refine_enter ?(unsafe = true) f =
else Pretype_errors.error_occur_check env sigma self c
in
(** Proceed to the refinement *)
+ let c = EConstr.Unsafe.to_constr c in
let sigma = match evkmain with
| None -> Evd.define self c sigma
| Some evk ->
diff --git a/proofs/refine.mli b/proofs/refine.mli
index a44632eff..1a254d578 100644
--- a/proofs/refine.mli
+++ b/proofs/refine.mli
@@ -21,7 +21,7 @@ val pr_constr :
(** {7 Refinement primitives} *)
-val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
+val refine : ?unsafe:bool -> EConstr.t Sigma.run -> unit tactic
(** In [refine ?unsafe t], [t] is a term with holes under some
[evar_map] context. The term [t] is used as a partial solution
for the current goal (refine is a goal-dependent tactic), the
@@ -30,17 +30,17 @@ val refine : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
tactic failures. If [unsafe] is [false] (default is [true]) [t] is
type-checked beforehand. *)
-val refine_one : ?unsafe:bool -> ('a * Constr.t) Sigma.run -> 'a tactic
+val refine_one : ?unsafe:bool -> ('a * EConstr.t) Sigma.run -> 'a tactic
(** A generalization of [refine] which assumes exactly one goal under focus *)
(** {7 Helper functions} *)
val with_type : Environ.env -> Evd.evar_map ->
- Term.constr -> Term.types -> Evd.evar_map * Term.constr
+ EConstr.constr -> EConstr.types -> Evd.evar_map * EConstr.constr
(** [with_type env sigma c t] ensures that [c] is of type [t]
inserting a coercion if needed. *)
-val refine_casted : ?unsafe:bool -> Constr.t Sigma.run -> unit tactic
+val refine_casted : ?unsafe:bool -> EConstr.t Sigma.run -> unit tactic
(** Like {!refine} except the refined term is coerced to the conclusion of the
current goal. *)
diff --git a/proofs/refiner.ml b/proofs/refiner.ml
index 9a0b56b84..5c7659ac0 100644
--- a/proofs/refiner.ml
+++ b/proofs/refiner.ml
@@ -22,7 +22,7 @@ let project x = x.sigma
(* Getting env *)
let pf_env gls = Global.env_of_context (Goal.V82.hyps (project gls) (sig_it gls))
-let pf_hyps gls = named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls))
+let pf_hyps gls = EConstr.named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls))
let refiner pr goal_sigma =
let (sgl,sigma') = prim_refiner pr goal_sigma.sigma goal_sigma.it in
@@ -198,10 +198,10 @@ let tclNOTSAMEGOAL (tac : tactic) goal =
destruct), this is not detected by this tactical. *)
let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma)
:Proof_type.goal list Evd.sigma =
- let oldhyps:Context.Named.t = pf_hyps goal in
+ let oldhyps = pf_hyps goal in
let rslt:Proof_type.goal list Evd.sigma = tac goal in
let { it = gls; sigma = sigma; } = rslt in
- let hyps:Context.Named.t list =
+ let hyps =
List.map (fun gl -> pf_hyps { it = gl; sigma=sigma; }) gls in
let cmp d1 d2 = Names.Id.equal (NamedDecl.get_id d1) (NamedDecl.get_id d2) in
let newhyps =
diff --git a/proofs/refiner.mli b/proofs/refiner.mli
index 6dcafb77a..56f5facf8 100644
--- a/proofs/refiner.mli
+++ b/proofs/refiner.mli
@@ -10,6 +10,7 @@
open Evd
open Proof_type
+open EConstr
(** The refiner (handles primitive rules and high-level tactics). *)
@@ -17,7 +18,7 @@ val sig_it : 'a sigma -> 'a
val project : 'a sigma -> evar_map
val pf_env : goal sigma -> Environ.env
-val pf_hyps : goal sigma -> Context.Named.t
+val pf_hyps : goal sigma -> named_context
val unpackage : 'a sigma -> evar_map ref * 'a
val repackage : evar_map ref -> 'a -> 'a sigma
diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml
index 030a3cbfb..b55f8ef11 100644
--- a/proofs/tacmach.ml
+++ b/proofs/tacmach.ml
@@ -41,11 +41,14 @@ let project = Refiner.project
let pf_env = Refiner.pf_env
let pf_hyps = Refiner.pf_hyps
+let test_conversion env sigma pb c1 c2 =
+ Reductionops.check_conv ~pb env sigma c1 c2
+
let pf_concl gls = Goal.V82.concl (project gls) (sig_it gls)
let pf_hyps_types gls =
let sign = Environ.named_context (pf_env gls) in
List.map (function LocalAssum (id,x)
- | LocalDef (id,_,x) -> id, x)
+ | LocalDef (id,_,x) -> id, EConstr.of_constr x)
sign
let pf_nth_hyp_id gls n = List.nth (pf_hyps gls) (n-1) |> NamedDecl.get_id
@@ -72,7 +75,7 @@ let pf_get_new_ids ids gls =
(fun id acc -> (next_ident_away id (acc@avoid))::acc)
ids []
-let pf_global gls id = Constrintern.construct_reference (pf_hyps gls) id
+let pf_global gls id = EConstr.of_constr (Constrintern.construct_reference (pf_hyps gls) id)
let pf_reduction_of_red_expr gls re c =
let (redfun, _) = reduction_of_red_expr (pf_env gls) re in
@@ -98,15 +101,15 @@ let pf_get_type_of = pf_reduce Retyping.get_type_of
let pf_conv_x gl = pf_reduce test_conversion gl Reduction.CONV
let pf_conv_x_leq gl = pf_reduce test_conversion gl Reduction.CUMUL
-let pf_const_value = pf_reduce (fun env _ -> constant_value_in env)
+let pf_const_value = pf_reduce (fun env _ c -> EConstr.of_constr (constant_value_in env c))
let pf_reduce_to_quantified_ind = pf_reduce reduce_to_quantified_ind
let pf_reduce_to_atomic_ind = pf_reduce reduce_to_atomic_ind
let pf_hnf_type_of gls = pf_get_type_of gls %> pf_whd_all gls
-let pf_is_matching = pf_apply Constr_matching.is_matching_conv
-let pf_matches = pf_apply Constr_matching.matches_conv
+let pf_is_matching gl p c = pf_apply Constr_matching.is_matching_conv gl p c
+let pf_matches gl p c = pf_apply Constr_matching.matches_conv gl p c
(********************************************)
(* Definition of the most primitive tactics *)
@@ -115,12 +118,15 @@ let pf_matches = pf_apply Constr_matching.matches_conv
let refiner = refiner
let internal_cut_no_check replace id t gl =
+ let t = EConstr.Unsafe.to_constr t in
refiner (Cut (true,replace,id,t)) gl
let internal_cut_rev_no_check replace id t gl =
+ let t = EConstr.Unsafe.to_constr t in
refiner (Cut (false,replace,id,t)) gl
let refine_no_check c gl =
+ let c = EConstr.Unsafe.to_constr c in
refiner (Refine c) gl
(* Versions with consistency checks *)
@@ -136,7 +142,7 @@ open Pp
let db_pr_goal sigma g =
let env = Goal.V82.env sigma g in
let penv = print_named_context env in
- let pc = print_constr_env env (Goal.V82.concl sigma g) in
+ let pc = print_constr_env env sigma (Goal.V82.concl sigma g) in
str" " ++ hv 0 (penv ++ fnl () ++
str "============================" ++ fnl () ++
str" " ++ pc) ++ fnl ()
@@ -165,7 +171,7 @@ module New = struct
(** We only check for the existence of an [id] in [hyps] *)
let gl = Proofview.Goal.assume gl in
let hyps = Proofview.Goal.hyps gl in
- Constrintern.construct_reference hyps id
+ EConstr.of_constr (Constrintern.construct_reference hyps id)
let pf_env = Proofview.Goal.env
let pf_concl = Proofview.Goal.concl
@@ -194,7 +200,7 @@ module New = struct
let pf_get_hyp id gl =
let hyps = Proofview.Goal.env gl in
let sign =
- try Environ.lookup_named id hyps
+ try EConstr.lookup_named id hyps
with Not_found -> raise (RefinerError (NoSuchHyp id))
in
sign
@@ -206,7 +212,7 @@ module New = struct
let env = Proofview.Goal.env gl in
let sign = Environ.named_context env in
List.map (function LocalAssum (id,x)
- | LocalDef (id,_,x) -> id, x)
+ | LocalDef (id,_,x) -> id, EConstr.of_constr x)
sign
let pf_last_hyp gl =
diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli
index 59f296f64..627a8e0e7 100644
--- a/proofs/tacmach.mli
+++ b/proofs/tacmach.mli
@@ -9,6 +9,7 @@
open Names
open Term
open Environ
+open EConstr
open Evd
open Proof_type
open Redexpr
@@ -33,18 +34,18 @@ val apply_sig_tac :
val pf_concl : goal sigma -> types
val pf_env : goal sigma -> env
-val pf_hyps : goal sigma -> Context.Named.t
+val pf_hyps : goal sigma -> named_context
(*i val pf_untyped_hyps : goal sigma -> (Id.t * constr) list i*)
val pf_hyps_types : goal sigma -> (Id.t * types) list
val pf_nth_hyp_id : goal sigma -> int -> Id.t
-val pf_last_hyp : goal sigma -> Context.Named.Declaration.t
+val pf_last_hyp : goal sigma -> named_declaration
val pf_ids_of_hyps : goal sigma -> Id.t list
val pf_global : goal sigma -> Id.t -> constr
val pf_unsafe_type_of : goal sigma -> constr -> types
val pf_type_of : goal sigma -> constr -> evar_map * types
val pf_hnf_type_of : goal sigma -> constr -> types
-val pf_get_hyp : goal sigma -> Id.t -> Context.Named.Declaration.t
+val pf_get_hyp : goal sigma -> Id.t -> named_declaration
val pf_get_hyp_typ : goal sigma -> Id.t -> types
val pf_get_new_id : Id.t -> goal sigma -> Id.t
@@ -67,8 +68,8 @@ val pf_whd_all : goal sigma -> constr -> constr
val pf_hnf_constr : goal sigma -> constr -> constr
val pf_nf : goal sigma -> constr -> constr
val pf_nf_betaiota : goal sigma -> constr -> constr
-val pf_reduce_to_quantified_ind : goal sigma -> types -> pinductive * types
-val pf_reduce_to_atomic_ind : goal sigma -> types -> pinductive * types
+val pf_reduce_to_quantified_ind : goal sigma -> types -> (inductive * EInstance.t) * types
+val pf_reduce_to_atomic_ind : goal sigma -> types -> (inductive * EInstance.t) * types
val pf_compute : goal sigma -> constr -> constr
val pf_unfoldn : (occurrences * evaluable_global_reference) list
-> goal sigma -> constr -> constr
@@ -106,31 +107,31 @@ module New : sig
val project : ('a, 'r) Proofview.Goal.t -> Evd.evar_map
val pf_env : ('a, 'r) Proofview.Goal.t -> Environ.env
- val pf_concl : ([ `NF ], 'r) Proofview.Goal.t -> types
+ val pf_concl : ('a, 'r) Proofview.Goal.t -> types
(** WRONG: To be avoided at all costs, it typechecks the term entirely but
forgets the universe constraints necessary to retypecheck it *)
- val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types
+ val pf_unsafe_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types
(** This function does no type inference and expects an already well-typed term.
It recomputes its type in the fastest way possible (no conversion is ever involved) *)
- val pf_get_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.types
+ val pf_get_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types
(** This function entirely type-checks the term and computes its type
and the implied universe constraints. *)
- val pf_type_of : ('a, 'r) Proofview.Goal.t -> Term.constr -> evar_map * Term.types
- val pf_conv_x : ('a, 'r) Proofview.Goal.t -> Term.constr -> Term.constr -> bool
+ val pf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> evar_map * types
+ val pf_conv_x : ('a, 'r) Proofview.Goal.t -> t -> t -> bool
- val pf_get_new_id : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> identifier
+ val pf_get_new_id : identifier -> ('a, 'r) Proofview.Goal.t -> identifier
val pf_ids_of_hyps : ('a, 'r) Proofview.Goal.t -> identifier list
val pf_hyps_types : ('a, 'r) Proofview.Goal.t -> (identifier * types) list
- val pf_get_hyp : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t
- val pf_get_hyp_typ : identifier -> ([ `NF ], 'r) Proofview.Goal.t -> types
- val pf_last_hyp : ([ `NF ], 'r) Proofview.Goal.t -> Context.Named.Declaration.t
+ val pf_get_hyp : identifier -> ('a, 'r) Proofview.Goal.t -> named_declaration
+ val pf_get_hyp_typ : identifier -> ('a, 'r) Proofview.Goal.t -> types
+ val pf_last_hyp : ('a, 'r) Proofview.Goal.t -> named_declaration
val pf_nf_concl : ([ `LZ ], 'r) Proofview.Goal.t -> types
- val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> pinductive * types
+ val pf_reduce_to_quantified_ind : ('a, 'r) Proofview.Goal.t -> types -> (inductive * EInstance.t) * types
val pf_hnf_constr : ('a, 'r) Proofview.Goal.t -> constr -> types
val pf_hnf_type_of : ('a, 'r) Proofview.Goal.t -> constr -> types
diff --git a/stm/stm.ml b/stm/stm.ml
index b0ad3f879..2b6ee5511 100644
--- a/stm/stm.ml
+++ b/stm/stm.ml
@@ -1800,9 +1800,10 @@ end = struct (* {{{ *)
Future.purify (fun () ->
let _,_,_,_,sigma0 = Proof.proof (Proof_global.give_me_the_proof ()) in
let g = Evd.find sigma0 r_goal in
+ let is_ground c = Evarutil.is_ground_term sigma0 (EConstr.of_constr c) in
if not (
- Evarutil.is_ground_term sigma0 Evd.(evar_concl g) &&
- List.for_all (Context.Named.Declaration.for_all (Evarutil.is_ground_term sigma0))
+ is_ground Evd.(evar_concl g) &&
+ List.for_all (Context.Named.Declaration.for_all is_ground)
Evd.(evar_context g))
then
CErrors.user_err ~hdr:"STM" (strbrk("the par: goal selector supports ground "^
@@ -1815,8 +1816,10 @@ end = struct (* {{{ *)
match Evd.(evar_body (find sigma r_goal)) with
| Evd.Evar_empty -> RespNoProgress
| Evd.Evar_defined t ->
+ let t = EConstr.of_constr t in
let t = Evarutil.nf_evar sigma t in
if Evarutil.is_ground_term sigma t then
+ let t = EConstr.Unsafe.to_constr t in
RespBuiltSubProof (t, Evd.evar_universe_context sigma)
else CErrors.user_err ~hdr:"STM" (str"The solution is not ground")
end) ()
@@ -1885,10 +1888,10 @@ end = struct (* {{{ *)
stm_pperr_endline (fun () -> hov 0 (
str"g=" ++ int (Evar.repr gid) ++ spc () ++
str"t=" ++ (Printer.pr_constr pt) ++ spc () ++
- str"uc=" ++ Evd.pr_evar_universe_context uc));
+ str"uc=" ++ Termops.pr_evar_universe_context uc));
(if abstract then Tactics.tclABSTRACT None else (fun x -> x))
(V82.tactic (Refiner.tclPUSHEVARUNIVCONTEXT uc) <*>
- Tactics.exact_no_check pt)
+ Tactics.exact_no_check (EConstr.of_constr pt))
with TacTask.NoProgress ->
if solve then Tacticals.New.tclSOLVE [] else tclUNIT ()
})
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 7558a707e..74cb7a364 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -6,14 +6,15 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*
-*)
+module CVars = Vars
+
open Pp
open Util
open CErrors
open Names
open Vars
open Termops
+open EConstr
open Environ
open Tacmach
open Genredexpr
@@ -83,25 +84,26 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl =
if poly then
(** Refresh the instance of the hint *)
let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in
- let map c = Vars.subst_univs_level_constr subst c in
+ let map c = CVars.subst_univs_level_constr subst c in
+ let emap c = Vars.subst_univs_level_constr subst c in
let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in
(** Only metas are mentioning the old universes. *)
let clenv = {
- templval = Evd.map_fl map clenv.templval;
- templtyp = Evd.map_fl map clenv.templtyp;
+ templval = Evd.map_fl emap clenv.templval;
+ templtyp = Evd.map_fl emap clenv.templtyp;
evd = Evd.map_metas map evd;
env = Proofview.Goal.env gl;
} in
- clenv, map c
+ clenv, emap c
else
let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in
{ clenv with evd = evd ; env = Proofview.Goal.env gl }, c
in clenv, c
let unify_resolve poly flags ((c : raw_hint), clenv) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let clenv, c = connect_hint_clenv poly c clenv gl in
- let clenv = Tacmach.New.of_old (fun gl -> clenv_unique_resolver ~flags clenv gl) gl in
+ let clenv = clenv_unique_resolver ~flags clenv gl in
Clenvtac.clenv_refine false clenv
end }
@@ -298,13 +300,13 @@ let flags_of_state st =
let auto_flags_of_state st =
auto_unif_flags_of full_transparent_state st false
-let hintmap_of secvars hdc concl =
+let hintmap_of sigma secvars hdc concl =
match hdc with
| None -> Hint_db.map_none ~secvars
| Some hdc ->
- if occur_existential concl then
- Hint_db.map_existential ~secvars hdc concl
- else Hint_db.map_auto ~secvars hdc concl
+ if occur_existential sigma concl then
+ Hint_db.map_existential sigma ~secvars hdc concl
+ else Hint_db.map_auto sigma ~secvars hdc concl
let exists_evaluable_reference env = function
| EvalConstRef _ -> true
@@ -328,25 +330,26 @@ let rec trivial_fail_db dbg mod_delta db_list local_db =
end })
in
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let sigma = Tacmach.New.project gl in
let secvars = compute_secvars gl in
Tacticals.New.tclFIRST
((dbg_assumption dbg)::intro_tac::
(List.map Tacticals.New.tclCOMPLETE
- (trivial_resolve dbg mod_delta db_list local_db secvars concl)))
+ (trivial_resolve sigma dbg mod_delta db_list local_db secvars concl)))
end }
-and my_find_search_nodelta db_list local_db secvars hdc concl =
+and my_find_search_nodelta sigma db_list local_db secvars hdc concl =
List.map (fun hint -> (None,hint))
- (List.map_append (hintmap_of secvars hdc concl) (local_db::db_list))
+ (List.map_append (hintmap_of sigma secvars hdc concl) (local_db::db_list))
and my_find_search mod_delta =
if mod_delta then my_find_search_delta
else my_find_search_nodelta
-and my_find_search_delta db_list local_db secvars hdc concl =
- let f = hintmap_of secvars hdc concl in
- if occur_existential concl then
+and my_find_search_delta sigma db_list local_db secvars hdc concl =
+ let f = hintmap_of sigma secvars hdc concl in
+ if occur_existential sigma concl then
List.map_append
(fun db ->
if Hint_db.use_dn db then
@@ -368,8 +371,8 @@ and my_find_search_delta db_list local_db secvars hdc concl =
match hdc with None -> Hint_db.map_none ~secvars db
| Some hdc ->
if (Id.Pred.is_empty ids && Cpred.is_empty csts)
- then Hint_db.map_auto ~secvars hdc concl db
- else Hint_db.map_existential ~secvars hdc concl db
+ then Hint_db.map_auto sigma ~secvars hdc concl db
+ else Hint_db.map_existential sigma ~secvars hdc concl db
in auto_flags_of_state st, l
in List.map (fun x -> (Some flags,x)) l)
(local_db::db_list)
@@ -402,23 +405,23 @@ and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t;poly=poly;db=
in
tclLOG dbg pr_hint (run_hint t tactic)
-and trivial_resolve dbg mod_delta db_list local_db secvars cl =
+and trivial_resolve sigma dbg mod_delta db_list local_db secvars cl =
try
let head =
- try let hdconstr = decompose_app_bound cl in
+ try let hdconstr = decompose_app_bound sigma cl in
Some hdconstr
with Bound -> None
in
List.map (tac_of_hint dbg db_list local_db cl)
(priority
- (my_find_search mod_delta db_list local_db secvars head cl))
+ (my_find_search mod_delta sigma db_list local_db secvars head cl))
with Not_found -> []
(** The use of the "core" database can be de-activated by passing
"nocore" amongst the databases. *)
let trivial ?(debug=Off) lems dbnames =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = make_db_list dbnames in
@@ -429,7 +432,7 @@ let trivial ?(debug=Off) lems dbnames =
end }
let full_trivial ?(debug=Off) lems =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = current_pure_db () in
@@ -449,15 +452,15 @@ let h_trivial ?(debug=Off) lems l = gen_trivial ~debug lems l
(* The classical Auto tactic *)
(**************************************************************************)
-let possible_resolve dbg mod_delta db_list local_db secvars cl =
+let possible_resolve sigma dbg mod_delta db_list local_db secvars cl =
try
let head =
- try let hdconstr = decompose_app_bound cl in
+ try let hdconstr = decompose_app_bound sigma cl in
Some hdconstr
with Bound -> None
in
List.map (tac_of_hint dbg db_list local_db cl)
- (my_find_search mod_delta db_list local_db secvars head cl)
+ (my_find_search mod_delta sigma db_list local_db secvars head cl)
with Not_found -> []
let extend_local_db decl db gl =
@@ -487,13 +490,14 @@ let search d n mod_delta db_list local_db =
Tacticals.New.tclORELSE0 (dbg_assumption d)
(Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db)
( Proofview.Goal.enter { enter = begin fun gl ->
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let sigma = Tacmach.New.project gl in
let secvars = compute_secvars gl in
let d' = incr_dbg d in
Tacticals.New.tclFIRST
(List.map
(fun ntac -> Tacticals.New.tclTHEN ntac (search d' (n-1) local_db))
- (possible_resolve d mod_delta db_list local_db secvars concl))
+ (possible_resolve sigma d mod_delta db_list local_db secvars concl))
end }))
end []
in
@@ -502,7 +506,7 @@ let search d n mod_delta db_list local_db =
let default_search_depth = ref 5
let delta_auto debug mod_delta n lems dbnames =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = make_db_list dbnames in
@@ -525,7 +529,7 @@ let new_auto ?(debug=Off) n = delta_auto debug true n
let default_auto = auto !default_search_depth [] []
let delta_full_auto ?(debug=Off) mod_delta n lems =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = current_pure_db () in
diff --git a/tactics/auto.mli b/tactics/auto.mli
index de0dbd483..32710e347 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -10,6 +10,7 @@
open Names
open Term
+open EConstr
open Clenv
open Pattern
open Decl_kinds
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index dae1cc9f1..e58ec5a31 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -92,7 +92,7 @@ type raw_rew_rule = Loc.t * constr Univ.in_universe_context_set * bool * Genarg.
let one_base general_rewrite_maybe_in tac_main bas =
let lrul = find_rewrites bas in
let try_rewrite dir ctx c tc =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in
let c' = Vars.subst_univs_level_constr subst c in
@@ -122,7 +122,7 @@ let autorewrite ?(conds=Naive) tac_main lbas =
Tacticals.New.tclTHEN tac
(one_base (fun dir c tac ->
let tac = (tac, conds) in
- general_rewrite dir AllOccurrences true false ~tac c)
+ general_rewrite dir AllOccurrences true false ~tac (EConstr.of_constr c))
tac_main bas))
(Proofview.tclUNIT()) lbas))
@@ -165,7 +165,7 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
| _ -> assert false) (* there must be at least an hypothesis *)
| _ -> assert false (* rewriting cannot complete a proof *)
in
- let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y z w) in
+ let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y (EConstr.of_constr z) w) in
Tacticals.New.tclMAP (fun id ->
Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS
(List.fold_left (fun tac bas ->
@@ -257,12 +257,12 @@ type hypinfo = {
let decompose_applied_relation metas env sigma c ctype left2right =
let find_rel ty =
- let eqclause = Clenv.mk_clenv_from_env env sigma None (c,ty) in
+ let eqclause = Clenv.mk_clenv_from_env env sigma None (EConstr.of_constr c,ty) in
let eqclause =
if metas then eqclause
else clenv_pose_metas_as_evars eqclause (Evd.undefined_metas eqclause.evd)
in
- let (equiv, args) = decompose_app (Clenv.clenv_type eqclause) in
+ let (equiv, args) = decompose_app (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in
let rec split_last_two = function
| [c1;c2] -> [],(c1, c2)
| x::y::z ->
@@ -272,11 +272,13 @@ let decompose_applied_relation metas env sigma c ctype left2right =
try
let others,(c1,c2) = split_last_two args in
let ty1, ty2 =
- Typing.unsafe_type_of env eqclause.evd c1, Typing.unsafe_type_of env eqclause.evd c2
+ Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c1), Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c2)
in
+ let ty = EConstr.Unsafe.to_constr ty in
+ let ty1 = EConstr.Unsafe.to_constr ty1 in
(* if not (evd_convertible env eqclause.evd ty1 ty2) then None *)
(* else *)
- Some { hyp_cl=eqclause; hyp_prf=(Clenv.clenv_value eqclause); hyp_ty = ty;
+ Some { hyp_cl=eqclause; hyp_prf=EConstr.Unsafe.to_constr (Clenv.clenv_value eqclause); hyp_ty = ty;
hyp_car=ty1; hyp_rel=mkApp (equiv, Array.of_list others);
hyp_l2r=left2right; hyp_left=c1; hyp_right=c2; }
with Not_found -> None
@@ -290,12 +292,12 @@ let decompose_applied_relation metas env sigma c ctype left2right =
| None -> None
let find_applied_relation metas loc env sigma c left2right =
- let ctype = Typing.unsafe_type_of env sigma c in
+ let ctype = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in
match decompose_applied_relation metas env sigma c ctype left2right with
| Some c -> c
| None ->
user_err ~loc ~hdr:"decompose_applied_relation"
- (str"The type" ++ spc () ++ Printer.pr_constr_env env sigma ctype ++
+ (str"The type" ++ spc () ++ Printer.pr_econstr_env env sigma ctype ++
spc () ++ str"of this term does not end with an applied relation.")
(* To add rewriting rules to a base *)
diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml
index 491bc8b4a..b4a235ba8 100644
--- a/tactics/btermdn.ml
+++ b/tactics/btermdn.ml
@@ -8,6 +8,7 @@
open Util
open Term
+open EConstr
open Names
open Pattern
open Globnames
@@ -38,18 +39,18 @@ let decomp_pat =
in
decrec []
-let decomp =
- let rec decrec acc c = match kind_of_term c with
+let decomp sigma t =
+ let rec decrec acc c = match EConstr.kind sigma c with
| App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f
| Proj (p, c) -> (mkConst (Projection.constant p), c :: acc)
| Cast (c1,_,_) -> decrec acc c1
| _ -> (c,acc)
in
- decrec []
+ decrec [] t
-let constr_val_discr t =
- let c, l = decomp t in
- match kind_of_term c with
+let constr_val_discr sigma t =
+ let c, l = decomp sigma t in
+ match EConstr.kind sigma c with
| Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l)
| Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l)
| Var id -> Label(GRLabel (VarRef id),l)
@@ -66,9 +67,9 @@ let constr_pat_discr t =
| PRef ((VarRef v) as ref), args -> Some(GRLabel ref,args)
| _ -> None
-let constr_val_discr_st (idpred,cpred) t =
- let c, l = decomp t in
- match kind_of_term c with
+let constr_val_discr_st sigma (idpred,cpred) t =
+ let c, l = decomp sigma t in
+ match EConstr.kind sigma c with
| Const (c,u) -> if Cpred.mem c cpred then Everything else Label(GRLabel (ConstRef c),l)
| Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l)
| Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l)
@@ -105,11 +106,11 @@ let bounded_constr_pat_discr_st st (t,depth) =
| None -> None
| Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
-let bounded_constr_val_discr_st st (t,depth) =
+let bounded_constr_val_discr_st sigma st (t,depth) =
if Int.equal depth 0 then
Nothing
else
- match constr_val_discr_st st t with
+ match constr_val_discr_st sigma st t with
| Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l)
| Nothing -> Nothing
| Everything -> Everything
@@ -122,11 +123,11 @@ let bounded_constr_pat_discr (t,depth) =
| None -> None
| Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
-let bounded_constr_val_discr (t,depth) =
+let bounded_constr_val_discr sigma (t,depth) =
if Int.equal depth 0 then
Nothing
else
- match constr_val_discr t with
+ match constr_val_discr sigma t with
| Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l)
| Nothing -> Nothing
| Everything -> Everything
@@ -162,13 +163,13 @@ struct
(fun dn (c,v) ->
Dn.rmv dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
- let lookup = function
+ let lookup sigma = function
| None ->
(fun dn t ->
- Dn.lookup dn bounded_constr_val_discr (t,!dnet_depth))
+ Dn.lookup dn (bounded_constr_val_discr sigma) (t,!dnet_depth))
| Some st ->
(fun dn t ->
- Dn.lookup dn (bounded_constr_val_discr_st st) (t,!dnet_depth))
+ Dn.lookup dn (bounded_constr_val_discr_st sigma st) (t,!dnet_depth))
let app f dn = Dn.app f dn
diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli
index 8ca5549b8..2a5e7c345 100644
--- a/tactics/btermdn.mli
+++ b/tactics/btermdn.mli
@@ -33,7 +33,7 @@ sig
val add : transparent_state option -> t -> (constr_pattern * Z.t) -> t
val rmv : transparent_state option -> t -> (constr_pattern * Z.t) -> t
- val lookup : transparent_state option -> t -> constr -> Z.t list
+ val lookup : Evd.evar_map -> transparent_state option -> t -> EConstr.constr -> Z.t list
val app : (Z.t -> unit) -> t -> unit
end
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index edfe21d34..ea1966093 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -17,6 +17,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Reduction
open Proof_type
open Tacticals
@@ -189,8 +190,7 @@ let set_typeclasses_strategy = function
| Bfs -> set_typeclasses_iterative_deepening true
let pr_ev evs ev =
- Printer.pr_constr_env (Goal.V82.env evs ev) evs
- (Evarutil.nf_evar evs (Goal.V82.concl evs ev))
+ Printer.pr_econstr_env (Goal.V82.env evs ev) evs (Goal.V82.concl evs ev)
(** Typeclasses instance search tactic / eauto *)
@@ -236,13 +236,13 @@ let e_give_exact flags poly (c,clenv) gl =
let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) ->
let clenv', c = connect_hint_clenv poly c clenv gls in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in
+ let clenv' = clenv_unique_resolver ~flags clenv' gls in
Clenvtac.clenv_refine true ~with_classes:false clenv'
end }
let unify_resolve poly flags = { enter = begin fun gls (c,_,clenv) ->
let clenv', _ = connect_hint_clenv poly c clenv gls in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in
+ let clenv' = clenv_unique_resolver ~flags clenv' gls in
Clenvtac.clenv_refine false ~with_classes:false clenv'
end }
@@ -264,7 +264,7 @@ let unify_resolve_refine poly flags gls ((c, t, ctx),n,clenv) =
sigma, c, t
in
let sigma', cl = Clenv.make_evar_clause env sigma ?len:n ty in
- let term = applistc term (List.map (fun x -> x.hole_evar) cl.cl_holes) in
+ let term = applist (term, List.map (fun x -> x.hole_evar) cl.cl_holes) in
let sigma' =
Evarconv.the_conv_x_leq env ~ts:flags.core_unify_flags.modulo_delta
cl.cl_concl concl sigma'
@@ -288,24 +288,24 @@ let clenv_of_prods poly nprods (c, clenv) gl =
let (c, _, _) = c in
if poly || Int.equal nprods 0 then Some (None, clenv)
else
- let ty = Retyping.get_type_of (Proofview.Goal.env gl)
- (Sigma.to_evar_map (Proofview.Goal.sigma gl)) c in
- let diff = nb_prod ty - nprods in
+ let sigma = Tacmach.New.project gl in
+ let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in
+ let diff = nb_prod sigma ty - nprods in
if Pervasives.(>=) diff 0 then
(* Was Some clenv... *)
Some (Some diff,
- Tacmach.New.of_old (fun gls -> mk_clenv_from_n gls (Some diff) (c,ty)) gl)
+ mk_clenv_from_n gl (Some diff) (c,ty))
else None
let with_prods nprods poly (c, clenv) f =
if get_typeclasses_limit_intros () then
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
try match clenv_of_prods poly nprods (c, clenv) gl with
| None -> Tacticals.New.tclZEROMSG (str"Not enough premisses")
| Some (diff, clenv') -> f.enter gl (c, diff, clenv')
with e when CErrors.noncritical e ->
Tacticals.New.tclZEROMSG (CErrors.print e) end }
- else Proofview.Goal.nf_enter
+ else Proofview.Goal.enter
{ enter = begin fun gl ->
if Int.equal nprods 0 then f.enter gl (c, None, clenv)
else Tacticals.New.tclZEROMSG (str"Not enough premisses") end }
@@ -351,21 +351,21 @@ let shelve_dependencies gls =
Feedback.msg_debug (str" shelving dependent subgoals: " ++ pr_gls sigma gls);
shelve_goals gls)
-let hintmap_of hdc secvars concl =
+let hintmap_of sigma hdc secvars concl =
match hdc with
| None -> fun db -> Hint_db.map_none secvars db
| Some hdc ->
fun db ->
if Hint_db.use_dn db then (* Using dnet *)
- Hint_db.map_eauto secvars hdc concl db
- else Hint_db.map_existential secvars hdc concl db
+ Hint_db.map_eauto sigma secvars hdc concl db
+ else Hint_db.map_existential sigma secvars hdc concl db
(** Hack to properly solve dependent evars that are typeclasses *)
let rec e_trivial_fail_db only_classes db_list local_db secvars =
let open Tacticals.New in
let open Tacmach.New in
let trivial_fail =
- Proofview.Goal.nf_enter { enter =
+ Proofview.Goal.enter { enter =
begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
@@ -376,7 +376,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars =
end }
in
let trivial_resolve =
- Proofview.Goal.nf_enter { enter =
+ Proofview.Goal.enter { enter =
begin fun gl ->
let tacs = e_trivial_resolve db_list local_db secvars only_classes
(project gl) (pf_concl gl) in
@@ -391,7 +391,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars =
and e_my_find_search db_list local_db secvars hdc complete only_classes sigma concl =
let open Proofview.Notations in
- let prods, concl = decompose_prod_assum concl in
+ let prods, concl = EConstr.decompose_prod_assum sigma concl in
let nprods = List.length prods in
let freeze =
try
@@ -399,12 +399,12 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co
| Some (hd,_) when only_classes ->
let cl = Typeclasses.class_info hd in
if cl.cl_strict then
- Evd.evars_of_term concl
+ Evarutil.undefined_evars_of_term sigma concl
else Evar.Set.empty
| _ -> Evar.Set.empty
with e when CErrors.noncritical e -> Evar.Set.empty
in
- let hint_of_db = hintmap_of hdc secvars concl in
+ let hint_of_db = hintmap_of sigma hdc secvars concl in
let hintl =
List.map_append
(fun db ->
@@ -480,13 +480,13 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co
in List.map tac_of_hint hintl
and e_trivial_resolve db_list local_db secvars only_classes sigma concl =
- let hd = try Some (decompose_app_bound concl) with Bound -> None in
+ let hd = try Some (decompose_app_bound sigma concl) with Bound -> None in
try
e_my_find_search db_list local_db secvars hd true only_classes sigma concl
with Not_found -> []
let e_possible_resolve db_list local_db secvars only_classes sigma concl =
- let hd = try Some (decompose_app_bound concl) with Bound -> None in
+ let hd = try Some (decompose_app_bound sigma concl) with Bound -> None in
try
e_my_find_search db_list local_db secvars hd false only_classes sigma concl
with Not_found -> []
@@ -511,13 +511,17 @@ let pr_depth l = prlist_with_sep (next_sep debug_seps) int (List.rev l)
let is_Prop env sigma concl =
let ty = Retyping.get_type_of env sigma concl in
- match kind_of_term ty with
- | Sort (Prop Null) -> true
+ match EConstr.kind sigma ty with
+ | Sort s ->
+ begin match ESorts.kind sigma s with
+ | Prop Null -> true
+ | _ -> false
+ end
| _ -> false
-let is_unique env concl =
+let is_unique env sigma concl =
try
- let (cl,u), args = dest_class_app env concl in
+ let (cl,u), args = dest_class_app env sigma concl in
cl.cl_unique
with e when CErrors.noncritical e -> false
@@ -560,14 +564,14 @@ let make_resolve_hyp env sigma st flags only_classes pri decl =
let id = NamedDecl.get_id decl in
let cty = Evarutil.nf_evar sigma (NamedDecl.get_type decl) in
let rec iscl env ty =
- let ctx, ar = decompose_prod_assum ty in
- match kind_of_term (fst (decompose_app ar)) with
+ let ctx, ar = decompose_prod_assum sigma ty in
+ match EConstr.kind sigma (fst (decompose_app sigma ar)) with
| Const (c,_) -> is_class (ConstRef c)
| Ind (i,_) -> is_class (IndRef i)
| _ ->
- let env' = Environ.push_rel_context ctx env in
- let ty' = whd_all env' ar in
- if not (Term.eq_constr ty' ar) then iscl env' ty'
+ let env' = push_rel_context ctx env in
+ let ty' = Reductionops.whd_all env' sigma ar in
+ if not (EConstr.eq_constr sigma ty' ar) then iscl env' ty'
else false
in
let is_class = iscl env cty in
@@ -587,7 +591,7 @@ let make_resolve_hyp env sigma st flags only_classes pri decl =
in
make_resolves env sigma ~name:(PathHints path)
(true,false,Flags.is_verbose()) info false
- (IsConstr (c,Univ.ContextSet.empty)))
+ (IsConstr (EConstr.of_constr c,Univ.ContextSet.empty)))
hints)
else []
in
@@ -605,7 +609,7 @@ let make_hints g st only_classes sign =
let consider =
try let t = hyp |> NamedDecl.get_id |> Global.lookup_named |> NamedDecl.get_type in
(* Section variable, reindex only if the type changed *)
- not (Term.eq_constr t (NamedDecl.get_type hyp))
+ not (EConstr.eq_constr (project g) (EConstr.of_constr t) (NamedDecl.get_type hyp))
with Not_found -> true
in
if consider then
@@ -660,7 +664,7 @@ module V85 = struct
then
cached_hints
else
- let hints = make_hints g st only_classes (Environ.named_context_of_val sign)
+ let hints = make_hints g st only_classes (EConstr.named_context_of_val sign)
in
cache := (only_classes, sign, hints); hints
@@ -677,7 +681,7 @@ module V85 = struct
let gls' =
List.map (fun g' ->
let env = Goal.V82.env s g' in
- let context = Environ.named_context_of_val (Goal.V82.hyps s g') in
+ let context = EConstr.named_context_of_val (Goal.V82.hyps s g') in
let hint = make_resolve_hyp env s (Hint_db.transparent_state info.hints)
(true,false,false) info.only_classes empty_hint_info (List.hd context) in
let ldb = Hint_db.add_list env s hint info.hints in
@@ -709,7 +713,7 @@ module V85 = struct
let needs_backtrack env evd oev concl =
if Option.is_empty oev || is_Prop env evd concl then
- occur_existential concl
+ occur_existential evd concl
else true
let hints_tac hints sk fk {it = gl,info; sigma = s} =
@@ -718,7 +722,7 @@ module V85 = struct
let tacgl = {it = gl; sigma = s;} in
let secvars = secvars_of_hyps (Environ.named_context_of_val (Goal.V82.hyps s gl)) in
let poss = e_possible_resolve hints info.hints secvars info.only_classes s concl in
- let unique = is_unique env concl in
+ let unique = is_unique env s concl in
let rec aux i foundone = function
| (tac, _, extern, name, pp) :: tl ->
let derivs = path_derivate info.auto_cut name in
@@ -783,10 +787,10 @@ module V85 = struct
let fk' =
(fun e ->
let do_backtrack =
- if unique then occur_existential concl
+ if unique then occur_existential tacgl.sigma concl
else if info.unique then true
else if List.is_empty gls' then
- needs_backtrack env s' info.is_evar concl
+ needs_backtrack env tacgl.sigma info.is_evar concl
else true
in
let e' = match foundone with None -> e
@@ -804,7 +808,7 @@ module V85 = struct
if foundone == None && !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.auto_depth ++ str": no match for " ++
- Printer.pr_constr_env (Goal.V82.env s gl) s concl ++
+ Printer.pr_econstr_env (Goal.V82.env s gl) s concl ++
spc () ++ str ", " ++ int (List.length poss) ++
str" possibilities");
match foundone with
@@ -987,13 +991,14 @@ module Search = struct
let sign = Goal.hyps g in
let (dir, onlyc, sign', cached_hints) = !autogoal_cache in
let cwd = Lib.cwd () in
+ let eq c1 c2 = EConstr.eq_constr (project g) c1 c2 in
if DirPath.equal cwd dir &&
(onlyc == only_classes) &&
- Context.Named.equal sign sign' &&
+ Context.Named.equal eq sign sign' &&
Hint_db.transparent_state cached_hints == st
then cached_hints
else
- let hints = make_hints {it = Goal.goal g; sigma = project g}
+ let hints = make_hints {it = Goal.goal (Proofview.Goal.assume g); sigma = project g}
st only_classes sign
in
autogoal_cache := (cwd, only_classes, sign, hints); hints
@@ -1025,7 +1030,7 @@ module Search = struct
NOT backtrack. *)
let needs_backtrack env evd unique concl =
if unique || is_Prop env evd concl then
- occur_existential concl
+ occur_existential evd concl
else true
let mark_unresolvables sigma goals =
@@ -1059,12 +1064,12 @@ module Search = struct
let concl = Goal.concl gl in
let sigma = Goal.sigma gl in
let s = Sigma.to_evar_map sigma in
- let unique = not info.search_dep || is_unique env concl in
+ let unique = not info.search_dep || is_unique env s concl in
let backtrack = needs_backtrack env s unique concl in
if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.search_depth ++ str": looking for " ++
- Printer.pr_constr_env (Goal.env gl) s concl ++
+ Printer.pr_econstr_env (Goal.env gl) s concl ++
(if backtrack then str" with backtracking"
else str" without backtracking"));
let secvars = compute_secvars gl in
@@ -1086,28 +1091,29 @@ module Search = struct
pr_depth (!idx :: info.search_depth) ++ str": " ++
Lazy.force pp ++
(if !foundone != true then
- str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl)
+ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl))
else mt ())
in
Feedback.msg_debug (msg ++ str " failed with " ++ CErrors.iprint ie)
else ()
in
- let tac_of gls i j = Goal.nf_enter { enter = fun gl' ->
+ let tac_of gls i j = Goal.enter { enter = fun gl' ->
let sigma' = Goal.sigma gl' in
let s' = Sigma.to_evar_map sigma' in
let _concl = Goal.concl gl' in
if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth (succ j :: i :: info.search_depth) ++ str" : " ++
- pr_ev s' (Proofview.Goal.goal gl'));
+ pr_ev s' (Proofview.Goal.goal (Proofview.Goal.assume gl')));
+ let eq c1 c2 = EConstr.eq_constr s' c1 c2 in
let hints' =
- if b && not (Context.Named.equal (Goal.hyps gl') (Goal.hyps gl))
+ if b && not (Context.Named.equal eq (Goal.hyps gl') (Goal.hyps gl))
then
let st = Hint_db.transparent_state info.search_hints in
make_autogoal_hints info.search_only_classes ~st gl'
else info.search_hints
in
- let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal gl') gls in
+ let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal (Proofview.Goal.assume gl')) gls in
let info' =
{ search_depth = succ j :: i :: info.search_depth;
last_tac = pp;
@@ -1124,7 +1130,7 @@ module Search = struct
(if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth (i :: info.search_depth) ++ str": " ++ Lazy.force pp
- ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl)
+ ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl))
++ str", " ++ int j ++ str" subgoal(s)" ++
(Option.cata (fun k -> str " in addition to the first " ++ int k)
(mt()) k)));
@@ -1138,7 +1144,7 @@ module Search = struct
try
let evi = Evd.find_undefined sigma ev in
if info.search_only_classes then
- Some (ev, not (is_class_type sigma (Evd.evar_concl evi)))
+ Some (ev, not (is_class_evar sigma evi))
else Some (ev, true)
with Not_found -> None
in
@@ -1194,7 +1200,7 @@ module Search = struct
if !foundone == false && !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.search_depth ++ str": no match for " ++
- Printer.pr_constr_env (Goal.env gl) s concl ++
+ Printer.pr_econstr_env (Goal.env gl) s concl ++
spc () ++ str ", " ++ int (List.length poss) ++
str" possibilities");
match e with
@@ -1205,7 +1211,7 @@ module Search = struct
else tclONCE (aux (NotApplicableEx,Exninfo.null) poss)
let hints_tac hints info kont : unit Proofview.tactic =
- Proofview.Goal.nf_enter
+ Proofview.Goal.enter
{ enter = fun gl -> hints_tac_gl hints info kont gl }
let intro_tac info kont gl =
@@ -1226,7 +1232,7 @@ module Search = struct
let intro info kont =
Proofview.tclBIND Tactics.intro
- (fun _ -> Proofview.Goal.nf_enter { enter = fun gl -> intro_tac info kont gl })
+ (fun _ -> Proofview.Goal.enter { enter = fun gl -> intro_tac info kont gl })
let rec search_tac hints limit depth =
let kont info =
@@ -1252,14 +1258,14 @@ module Search = struct
if false (* In 8.6, still allow non-class goals only_classes && not (is_class_type sigma (Goal.concl gl)) *) then
Tacticals.New.tclZEROMSG (str"Not a subgoal for a class")
else
- let dep = dep || Proofview.unifiable sigma (Goal.goal gl) gls in
+ let dep = dep || Proofview.unifiable sigma (Goal.goal (Proofview.Goal.assume gl)) gls in
let info = make_autogoal ?st only_classes dep (cut_of_hints hints) i gl in
search_tac hints depth 1 info
let search_tac ?(st=full_transparent_state) only_classes dep hints depth =
let open Proofview in
let tac sigma gls i =
- Goal.nf_enter
+ Goal.enter
{ enter = fun gl ->
search_tac_gl ~st only_classes dep hints depth (succ i) sigma gls gl }
in
@@ -1410,8 +1416,8 @@ module Intpart = Unionfind.Make(Evar.Set)(Evar.Map)
let deps_of_constraints cstrs evm p =
List.iter (fun (_, _, x, y) ->
- let evx = Evarutil.undefined_evars_of_term evm x in
- let evy = Evarutil.undefined_evars_of_term evm y in
+ let evx = Evarutil.undefined_evars_of_term evm (EConstr.of_constr x) in
+ let evy = Evarutil.undefined_evars_of_term evm (EConstr.of_constr y) in
Intpart.union_set (Evar.Set.union evx evy) p)
cstrs
@@ -1455,7 +1461,7 @@ let error_unresolvable env comp evd =
| Some s -> Evar.Set.mem ev s
in
let fold ev evi (found, accu) =
- let ev_class = class_of_constr evi.evar_concl in
+ let ev_class = class_of_constr evd (EConstr.of_constr evi.evar_concl) in
if not (Option.is_empty ev_class) && is_part ev then
(* focus on one instance if only one was searched for *)
if not found then (true, Some ev)
@@ -1560,9 +1566,10 @@ let _ =
Hook.set Typeclasses.solve_all_instances_hook solve_inst
let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique =
- let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in
+ let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env sigma gl in
let (gl,t,sigma) =
Goal.V82.mk_goal sigma nc gl Store.empty in
+ let (ev, _) = destEvar sigma t in
let gls = { it = gl ; sigma = sigma; } in
let hints = searchtable_map typeclasses_db in
let st = Hint_db.transparent_state hints in
@@ -1577,9 +1584,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique =
with Refiner.FailError _ -> raise Not_found
in
let evd = sig_sig gls' in
- let t' = let (ev, inst) = destEvar t in
- mkEvar (ev, Array.of_list subst)
- in
+ let t' = mkEvar (ev, Array.of_list subst) in
let term = Evarutil.nf_evar evd t' in
evd, term
@@ -1590,21 +1595,22 @@ let _ =
(** Take the head of the arity of a constr.
Used in the partial application tactic. *)
-let rec head_of_constr t =
- let t = strip_outer_cast(collapse_appl t) in
- match kind_of_term t with
- | Prod (_,_,c2) -> head_of_constr c2
- | LetIn (_,_,_,c2) -> head_of_constr c2
- | App (f,args) -> head_of_constr f
+let rec head_of_constr sigma t =
+ let t = strip_outer_cast sigma (collapse_appl sigma t) in
+ match EConstr.kind sigma t with
+ | Prod (_,_,c2) -> head_of_constr sigma c2
+ | LetIn (_,_,_,c2) -> head_of_constr sigma c2
+ | App (f,args) -> head_of_constr sigma f
| _ -> t
let head_of_constr h c =
- let c = head_of_constr c in
+ Proofview.tclEVARMAP >>= fun sigma ->
+ let c = head_of_constr sigma c in
letin_tac None (Name h) c None Locusops.allHyps
let not_evar c =
Proofview.tclEVARMAP >>= fun sigma ->
- match Evarutil.kind_of_term_upto sigma c with
+ match EConstr.kind sigma c with
| Evar _ -> Tacticals.New.tclFAIL 0 (str"Evar")
| _ -> Proofview.tclUNIT ()
@@ -1612,17 +1618,15 @@ let is_ground c gl =
if Evarutil.is_ground_term (project gl) c then tclIDTAC gl
else tclFAIL 0 (str"Not ground") gl
-let autoapply c i gl =
+let autoapply c i =
let open Proofview.Notations in
+ Proofview.Goal.enter { enter = begin fun gl ->
let flags = auto_unif_flags Evar.Set.empty
(Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in
- let cty = pf_unsafe_type_of gl c in
+ let cty = Tacmach.New.pf_unsafe_type_of gl c in
let ce = mk_clenv_from gl (c,cty) in
- let enter gl =
(unify_e_resolve false flags).enter gl
((c,cty,Univ.ContextSet.empty),0,ce) <*>
Proofview.tclEVARMAP >>= (fun sigma ->
let sigma = Typeclasses.mark_unresolvables ~filter:Typeclasses.all_goals sigma in
- Proofview.Unsafe.tclEVARS sigma)
- in
- Proofview.V82.of_tactic (Proofview.Goal.nf_enter { enter }) gl
+ Proofview.Unsafe.tclEVARS sigma) end }
diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli
index 76760db02..a38be5972 100644
--- a/tactics/class_tactics.mli
+++ b/tactics/class_tactics.mli
@@ -10,6 +10,7 @@
open Names
open Constr
+open EConstr
open Tacmach
val catchable : exn -> bool
@@ -28,13 +29,13 @@ val typeclasses_eauto : ?only_classes:bool -> ?st:transparent_state -> ?strategy
depth:(Int.t option) ->
Hints.hint_db_name list -> unit Proofview.tactic
-val head_of_constr : Id.t -> Term.constr -> unit Proofview.tactic
+val head_of_constr : Id.t -> constr -> unit Proofview.tactic
val not_evar : constr -> unit Proofview.tactic
val is_ground : constr -> tactic
-val autoapply : constr -> Hints.hint_db_name -> tactic
+val autoapply : constr -> Hints.hint_db_name -> unit Proofview.tactic
module Search : sig
val eauto_tac :
diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml
index 6b29f574c..63f923dfd 100644
--- a/tactics/contradiction.ml
+++ b/tactics/contradiction.ml
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Hipattern
open Tactics
open Coqlib
@@ -19,6 +20,7 @@ module NamedDecl = Context.Named.Declaration
(* Absurd *)
let mk_absurd_proof t =
+ let build_coq_not () = EConstr.of_constr (build_coq_not ()) in
let id = Namegen.default_dependent_ident in
mkLambda (Names.Name id,mkApp(build_coq_not (),[|t|]),
mkLambda (Names.Name id,t,mkApp (mkRel 2,[|mkRel 1|])))
@@ -33,7 +35,7 @@ let absurd c =
let t = j.Environ.utj_val in
let tac =
Tacticals.New.tclTHENLIST [
- elim_type (build_coq_False ());
+ elim_type (EConstr.of_constr (build_coq_False ()));
Simple.apply (mk_absurd_proof t)
] in
Sigma.Unsafe.of_pair (tac, sigma)
@@ -66,18 +68,18 @@ let contradiction_context =
let id = NamedDecl.get_id d in
let typ = nf_evar sigma (NamedDecl.get_type d) in
let typ = whd_all env sigma typ in
- if is_empty_type typ then
+ if is_empty_type sigma typ then
simplest_elim (mkVar id)
- else match kind_of_term typ with
- | Prod (na,t,u) when is_empty_type u ->
+ else match EConstr.kind sigma typ with
+ | Prod (na,t,u) when is_empty_type sigma u ->
let is_unit_or_eq =
- if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type t
+ if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma t
else None in
Tacticals.New.tclORELSE
(match is_unit_or_eq with
| Some _ ->
- let hd,args = decompose_app t in
- let (ind,_ as indu) = destInd hd in
+ let hd,args = decompose_app sigma t in
+ let (ind,_ as indu) = destInd sigma hd in
let nparams = Inductiveops.inductive_nparams_env env ind in
let params = Util.List.firstn nparams args in
let p = applist ((mkConstructUi (indu,1)), params) in
@@ -102,20 +104,19 @@ let contradiction_context =
end }
let is_negation_of env sigma typ t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (na,t,u) ->
- let u = nf_evar sigma u in
- is_empty_type u && is_conv_leq env sigma typ t
+ is_empty_type sigma u && is_conv_leq env sigma typ t
| _ -> false
let contradiction_term (c,lbind as cl) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
let type_of = Tacmach.New.pf_unsafe_type_of gl in
let typ = type_of c in
let _, ccl = splay_prod env sigma typ in
- if is_empty_type ccl then
+ if is_empty_type sigma ccl then
Tacticals.New.tclTHEN
(elim false None cl None)
(Tacticals.New.tclTRY assumption)
@@ -123,7 +124,7 @@ let contradiction_term (c,lbind as cl) =
Proofview.tclORELSE
begin
if lbind = NoBindings then
- filter_hyp (is_negation_of env sigma typ)
+ filter_hyp (fun c -> is_negation_of env sigma typ c)
(fun id -> simplest_elim (mkApp (mkVar id,[|c|])))
else
Proofview.tclZERO Not_found
diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli
index b876aee90..510b135b0 100644
--- a/tactics/contradiction.mli
+++ b/tactics/contradiction.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Misctypes
val absurd : constr -> unit Proofview.tactic
diff --git a/tactics/eauto.ml b/tactics/eauto.ml
index 885183174..8d1e0e507 100644
--- a/tactics/eauto.ml
+++ b/tactics/eauto.ml
@@ -13,6 +13,7 @@ open Names
open Nameops
open Term
open Termops
+open EConstr
open Proof_type
open Tacticals
open Tacmach
@@ -32,7 +33,8 @@ let e_give_exact ?(flags=eauto_unif_flags) c =
Proofview.Goal.enter { enter = begin fun gl ->
let t1 = Tacmach.New.pf_unsafe_type_of gl c in
let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in
- if occur_existential t1 || occur_existential t2 then
+ let sigma = Tacmach.New.project gl in
+ if occur_existential sigma t1 || occur_existential sigma t2 then
Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c)
else exact_check c
end }
@@ -86,7 +88,7 @@ let rec prolog l n gl =
let out_term = function
| IsConstr (c, _) -> c
- | IsGlobRef gr -> fst (Universes.fresh_global_instance (Global.env ()) gr)
+ | IsGlobRef gr -> EConstr.of_constr (fst (Universes.fresh_global_instance (Global.env ()) gr))
let prolog_tac l n =
Proofview.V82.tactic begin fun gl ->
@@ -110,22 +112,21 @@ open Auto
let priority l = List.map snd (List.filter (fun (pr,_) -> Int.equal pr 0) l)
let unify_e_resolve poly flags (c,clenv) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let clenv', c = connect_hint_clenv poly c clenv gl in
- Proofview.V82.tactic
- (fun gls ->
- let clenv' = clenv_unique_resolver ~flags clenv' gls in
- tclTHEN (Refiner.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd))
- (Proofview.V82.of_tactic (Tactics.Simple.eapply c)) gls)
+ let clenv' = clenv_unique_resolver ~flags clenv' gl in
+ Proofview.tclTHEN
+ (Proofview.Unsafe.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd))
+ (Tactics.Simple.eapply c)
end }
-let hintmap_of secvars hdc concl =
+let hintmap_of sigma secvars hdc concl =
match hdc with
| None -> fun db -> Hint_db.map_none ~secvars db
| Some hdc ->
- if occur_existential concl then
- (fun db -> Hint_db.map_existential ~secvars hdc concl db)
- else (fun db -> Hint_db.map_auto ~secvars hdc concl db)
+ if occur_existential sigma concl then
+ (fun db -> Hint_db.map_existential sigma ~secvars hdc concl db)
+ else (fun db -> Hint_db.map_auto sigma ~secvars hdc concl db)
(* FIXME: should be (Hint_db.map_eauto hdc concl db) *)
let e_exact poly flags (c,clenv) =
@@ -137,7 +138,7 @@ let e_exact poly flags (c,clenv) =
end }
let rec e_trivial_fail_db db_list local_db =
- let next = Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let next = Proofview.Goal.enter { enter = begin fun gl ->
let d = Tacmach.New.pf_last_hyp gl in
let hintl = make_resolve_hyp (Tacmach.New.pf_env gl) (Tacmach.New.project gl) d in
e_trivial_fail_db db_list (Hint_db.add_list (Tacmach.New.pf_env gl) (Tacmach.New.project gl) hintl local_db)
@@ -147,13 +148,13 @@ let rec e_trivial_fail_db db_list local_db =
let tacl =
registered_e_assumption ::
(Tacticals.New.tclTHEN Tactics.intro next) ::
- (List.map fst (e_trivial_resolve db_list local_db secvars (Tacmach.New.pf_nf_concl gl)))
+ (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_concl gl)))
in
Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl)
end }
-and e_my_find_search db_list local_db secvars hdc concl =
- let hint_of_db = hintmap_of secvars hdc concl in
+and e_my_find_search sigma db_list local_db secvars hdc concl =
+ let hint_of_db = hintmap_of sigma secvars hdc concl in
let hintl =
List.map_append (fun db ->
let flags = auto_flags_of_state (Hint_db.transparent_state db) in
@@ -181,15 +182,15 @@ and e_my_find_search db_list local_db secvars hdc concl =
in
List.map tac_of_hint hintl
-and e_trivial_resolve db_list local_db secvars gl =
- let hd = try Some (decompose_app_bound gl) with Bound -> None in
- try priority (e_my_find_search db_list local_db secvars hd gl)
+and e_trivial_resolve sigma db_list local_db secvars gl =
+ let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in
+ try priority (e_my_find_search sigma db_list local_db secvars hd gl)
with Not_found -> []
-let e_possible_resolve db_list local_db secvars gl =
- let hd = try Some (decompose_app_bound gl) with Bound -> None in
+let e_possible_resolve sigma db_list local_db secvars gl =
+ let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in
try List.map (fun (b, (tac, pp)) -> (tac, b, pp))
- (e_my_find_search db_list local_db secvars hd gl)
+ (e_my_find_search sigma db_list local_db secvars hd gl)
with Not_found -> []
let find_first_goal gls =
@@ -287,9 +288,9 @@ module SearchProblem = struct
in
let rec_tacs =
let l =
- let concl = Reductionops.nf_evar (project g)(pf_concl g) in
+ let concl = Reductionops.nf_evar (project g) (pf_concl g) in
filter_tactics s.tacres
- (e_possible_resolve s.dblist (List.hd s.localdb) secvars concl)
+ (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl)
in
List.map
(fun (lgls, cost, pp) ->
@@ -464,18 +465,19 @@ let autounfold_tac db cls =
in
autounfold dbs cls
-let unfold_head env (ids, csts) c =
+let unfold_head env sigma (ids, csts) c =
let rec aux c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id when Id.Set.mem id ids ->
(match Environ.named_body id env with
- | Some b -> true, b
+ | Some b -> true, EConstr.of_constr b
| None -> false, c)
- | Const (cst,u as c) when Cset.mem cst csts ->
- true, Environ.constant_value_in env c
+ | Const (cst, u) when Cset.mem cst csts ->
+ let u = EInstance.kind sigma u in
+ true, EConstr.of_constr (Environ.constant_value_in env (cst, u))
| App (f, args) ->
(match aux f with
- | true, f' -> true, Reductionops.whd_betaiota Evd.empty (mkApp (f', args))
+ | true, f' -> true, Reductionops.whd_betaiota sigma (mkApp (f', args))
| false, _ ->
let done_, args' =
Array.fold_left_i (fun i (done_, acc) arg ->
@@ -489,7 +491,7 @@ let unfold_head env (ids, csts) c =
else false, c)
| _ ->
let done_ = ref false in
- let c' = map_constr (fun c ->
+ let c' = EConstr.map sigma (fun c ->
if !done_ then c else
let x, c' = aux c in
done_ := x; c') c
@@ -497,8 +499,9 @@ let unfold_head env (ids, csts) c =
in aux c
let autounfold_one db cl =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
let st =
List.fold_left (fun (i,c) dbname ->
@@ -508,7 +511,7 @@ let autounfold_one db cl =
let (ids, csts) = Hint_db.unfolds db in
(Id.Set.union ids i, Cset.union csts c)) (Id.Set.empty, Cset.empty) db
in
- let did, c' = unfold_head env st
+ let did, c' = unfold_head env sigma st
(match cl with Some (id, _) -> Tacmach.New.pf_get_hyp_typ id gl | None -> concl)
in
if did then
diff --git a/tactics/eauto.mli b/tactics/eauto.mli
index 1f69e4ab3..e2006ac1e 100644
--- a/tactics/eauto.mli
+++ b/tactics/eauto.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Proof_type
open Hints
open Tactypes
diff --git a/tactics/elim.ml b/tactics/elim.ml
index 3f0c01a29..e37ec6bce 100644
--- a/tactics/elim.ml
+++ b/tactics/elim.ml
@@ -10,6 +10,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Inductiveops
open Hipattern
open Tacmach.New
@@ -79,11 +80,12 @@ let up_to_delta = ref false (* true *)
let general_decompose recognizer c =
Proofview.Goal.enter { enter = begin fun gl ->
let type_of = pf_unsafe_type_of gl in
+ let sigma = project gl in
let typc = type_of c in
tclTHENS (cut typc)
[ tclTHEN (intro_using tmphyp_name)
(onLastHypId
- (ifOnHyp recognizer (general_decompose_aux recognizer)
+ (ifOnHyp (recognizer sigma) (general_decompose_aux (recognizer sigma))
(fun id -> clear [id])));
exact_no_check c ]
end }
@@ -95,24 +97,24 @@ let head_in indl t gl =
let ity,_ =
if !up_to_delta
then find_mrectype env sigma t
- else extract_mrectype t
+ else extract_mrectype sigma t
in List.exists (fun i -> eq_ind (fst i) (fst ity)) indl
with Not_found -> false
let decompose_these c l =
Proofview.Goal.enter { enter = begin fun gl ->
let indl = List.map (fun x -> x, Univ.Instance.empty) l in
- general_decompose (fun (_,t) -> head_in indl t gl) c
+ general_decompose (fun sigma (_,t) -> head_in indl t gl) c
end }
let decompose_and c =
general_decompose
- (fun (_,t) -> is_record t)
+ (fun sigma (_,t) -> is_record sigma t)
c
let decompose_or c =
general_decompose
- (fun (_,t) -> is_disjunction t)
+ (fun sigma (_,t) -> is_disjunction sigma t)
c
let h_decompose l c = decompose_these c l
@@ -131,9 +133,9 @@ let induction_trailer abs_i abs_j bargs =
(tclDO (abs_j - abs_i) intro)
(onLastHypId
(fun id ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let idty = pf_unsafe_type_of gl (mkVar id) in
- let fvty = global_vars (pf_env gl) idty in
+ let fvty = global_vars (pf_env gl) (project gl) idty in
let possible_bring_hyps =
(List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums
in
@@ -153,7 +155,7 @@ let induction_trailer abs_i abs_j bargs =
))
let double_ind h1 h2 =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let abs_i = depth_of_quantified_hypothesis true h1 gl in
let abs_j = depth_of_quantified_hypothesis true h2 gl in
let abs =
diff --git a/tactics/elim.mli b/tactics/elim.mli
index 29c441463..dc1af79ba 100644
--- a/tactics/elim.mli
+++ b/tactics/elim.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Tacticals
open Misctypes
open Tactypes
diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml
index 1a67bedc2..bac3980d2 100644
--- a/tactics/eqdecide.ml
+++ b/tactics/eqdecide.ml
@@ -16,6 +16,7 @@ open Util
open Names
open Namegen
open Term
+open EConstr
open Declarations
open Tactics
open Tacticals.New
@@ -25,6 +26,7 @@ open Misctypes
open Tactypes
open Hipattern
open Pretyping
+open Proofview.Notations
open Tacmach.New
open Coqlib
@@ -50,7 +52,9 @@ open Coqlib
Eduardo Gimenez (30/3/98).
*)
-let clear_last = (onLastHyp (fun c -> (clear [destVar c])))
+let clear_last =
+ Proofview.tclEVARMAP >>= fun sigma ->
+ (onLastHyp (fun c -> (clear [destVar sigma c])))
let choose_eq eqonleft =
if eqonleft then
@@ -73,7 +77,7 @@ let mkBranches c1 c2 =
intros]
let discrHyp id =
- let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in
+ let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in
let tac c = Equality.discr_tac false (Some (None, ElimOnConstr c)) in
Tacticals.New.tclDELAYEDWITHHOLES false c tac
@@ -85,7 +89,9 @@ let solveNoteqBranch side =
(* Constructs the type {c1=c2}+{~c1=c2} *)
let make_eq () =
-(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ()))
+let build_coq_not () = EConstr.of_constr (build_coq_not ())
+let build_coq_sumbool () = EConstr.of_constr (build_coq_sumbool ())
let mkDecideEqGoal eqonleft op rectype c1 c2 =
let equality = mkApp(make_eq(), [|rectype; c1; c2|]) in
@@ -121,7 +127,7 @@ let eqCase tac =
tclTHEN intro (onLastHypId tac)
let injHyp id =
- let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in
+ let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in
let tac c = Equality.injClause None false (Some (None, ElimOnConstr c)) in
Tacticals.New.tclDELAYEDWITHHOLES false c tac
@@ -141,8 +147,8 @@ open Proofview.Notations
(* spiwack: a small wrapper around [Hipattern]. *)
-let match_eqdec c =
- try Proofview.tclUNIT (match_eqdec c)
+let match_eqdec sigma c =
+ try Proofview.tclUNIT (match_eqdec sigma c)
with PatternMatchingFailure -> Proofview.tclZERO PatternMatchingFailure
(* /spiwack *)
@@ -170,11 +176,12 @@ let solveEqBranch rectype =
Proofview.tclORELSE
begin
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
- match_eqdec concl >>= fun (eqonleft,op,lhs,rhs,_) ->
+ let concl = pf_concl gl in
+ let sigma = project gl in
+ match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) ->
let (mib,mip) = Global.lookup_inductive rectype in
let nparams = mib.mind_nparams in
- let getargs l = List.skipn nparams (snd (decompose_app l)) in
+ let getargs l = List.skipn nparams (snd (decompose_app sigma l)) in
let rargs = getargs rhs
and largs = getargs lhs in
solveArg [] eqonleft op largs rargs
@@ -187,7 +194,7 @@ let solveEqBranch rectype =
(* The tactic Decide Equality *)
-let hd_app c = match kind_of_term c with
+let hd_app sigma c = match EConstr.kind sigma c with
| App (h,_) -> h
| _ -> c
@@ -195,10 +202,11 @@ let decideGralEquality =
Proofview.tclORELSE
begin
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
- match_eqdec concl >>= fun (eqonleft,_,c1,c2,typ) ->
- let headtyp = hd_app (pf_compute gl typ) in
- begin match kind_of_term headtyp with
+ let concl = pf_concl gl in
+ let sigma = project gl in
+ match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) ->
+ let headtyp = hd_app sigma (pf_compute gl typ) in
+ begin match EConstr.kind sigma headtyp with
| Ind (mi,_) -> Proofview.tclUNIT mi
| _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.")
end >>= fun rectype ->
diff --git a/tactics/eqdecide.mli b/tactics/eqdecide.mli
index cb48a5bcc..dca1780b7 100644
--- a/tactics/eqdecide.mli
+++ b/tactics/eqdecide.mli
@@ -14,4 +14,4 @@
val decideEqualityGoal : unit Proofview.tactic
-val compare : Constr.t -> Constr.t -> unit Proofview.tactic
+val compare : EConstr.t -> EConstr.t -> unit Proofview.tactic
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index c94dcfa9d..b08456f2f 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -73,13 +73,27 @@ let build_dependent_inductive ind (mib,mip) =
let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in
applist
(mkIndU ind,
- Context.Rel.to_extended_list mip.mind_nrealdecls mib.mind_params_ctxt
- @ Context.Rel.to_extended_list 0 realargs)
+ Context.Rel.to_extended_list mkRel mip.mind_nrealdecls mib.mind_params_ctxt
+ @ Context.Rel.to_extended_list mkRel 0 realargs)
+
+let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na
+let name_assumption env = function
+| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
+| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+
+let name_context env hyps =
+ snd
+ (List.fold_left
+ (fun (env,hyps) d ->
+ let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
+ (env,[]) (List.rev hyps))
let my_it_mkLambda_or_LetIn s c = it_mkLambda_or_LetIn c s
-let my_it_mkProd_or_LetIn s c = it_mkProd_or_LetIn c s
+let my_it_mkProd_or_LetIn s c = Term.it_mkProd_or_LetIn c s
let my_it_mkLambda_or_LetIn_name s c =
- it_mkLambda_or_LetIn_name (Global.env()) c s
+ let env = Global.env () in
+ let mkLambda_or_LetIn_name d b = mkLambda_or_LetIn (name_assumption env d) b in
+ List.fold_left (fun c d -> mkLambda_or_LetIn_name d c) c s
let get_coq_eq ctx =
try
@@ -120,7 +134,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 eq_constr params2 constrargs) then
+ if not (List.equal Term.eq_constr 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)
@@ -172,7 +186,7 @@ let build_sym_scheme env ind =
let (mib,mip as specif),nrealargs,realsign,paramsctxt,paramsctxt1 =
get_sym_eq_data env indu in
let cstr n =
- mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in
+ mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
@@ -185,7 +199,7 @@ let build_sym_scheme env ind =
my_it_mkLambda_or_LetIn_name
(lift_rel_context (nrealargs+1) realsign_ind)
(mkApp (mkIndU indu,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect 1 nrealargs;
rel_vect (2*nrealargs+2) nrealargs])),
mkRel 1 (* varH *),
@@ -226,13 +240,13 @@ let build_sym_involutive_scheme env ind =
get_sym_eq_data env indu in
let eq,eqrefl,ctx = get_coq_eq ctx in
let sym, ctx, eff = const_of_scheme sym_scheme_kind env ind ctx in
- let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect n paramsctxt) in
+ let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect mkRel n paramsctxt) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_C =
mkApp
(mkIndU indu, Array.append
- (Context.Rel.to_extended_vect (nrealargs+1) mib.mind_params_ctxt)
+ (Context.Rel.to_extended_vect mkRel (nrealargs+1) mib.mind_params_ctxt)
(rel_vect (nrealargs+1) nrealargs)) in
let realsign_ind =
name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
@@ -246,15 +260,15 @@ let build_sym_involutive_scheme env ind =
(mkApp (eq,[|
mkApp
(mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect (2*nrealargs+2) nrealargs;
rel_vect 1 nrealargs]);
mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect 1 nrealargs;
rel_vect (2*nrealargs+2) nrealargs;
[|mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect (2*nrealargs+2) nrealargs;
rel_vect 1 nrealargs;
[|mkRel 1|]])|]]);
@@ -337,7 +351,7 @@ let build_l2r_rew_scheme dep env ind kind =
let eq,eqrefl,ctx = get_coq_eq ctx in
let cstr n p =
mkApp (mkConstructUi(indu,1),
- Array.concat [Context.Rel.to_extended_vect n paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -345,12 +359,12 @@ let build_l2r_rew_scheme dep env ind kind =
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_P =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect nrealargs nrealargs]) in
let applied_ind_G =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+3) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+3) paramsctxt1;
rel_vect (nrealargs+3) nrealargs;
rel_vect 0 nrealargs]) in
let realsign_P = lift_rel_context nrealargs realsign in
@@ -361,10 +375,10 @@ let build_l2r_rew_scheme dep env ind kind =
lift_rel_context (nrealargs+3) realsign) in
let applied_sym_C n =
mkApp(sym,
- Array.append (Context.Rel.to_extended_vect n mip.mind_arity_ctxt) [|mkVar varH|]) in
+ Array.append (Context.Rel.to_extended_vect mkRel n mip.mind_arity_ctxt) [|mkVar varH|]) in
let applied_sym_G =
mkApp(sym,
- Array.concat [Context.Rel.to_extended_vect (nrealargs*3+4) paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel (nrealargs*3+4) paramsctxt1;
rel_vect (nrealargs+4) nrealargs;
rel_vect 1 nrealargs;
[|mkRel 1|]]) in
@@ -374,7 +388,7 @@ let build_l2r_rew_scheme dep env ind kind =
let ci = make_case_info (Global.env()) ind RegularStyle in
let cieq = make_case_info (Global.env()) (fst (destInd eq)) RegularStyle in
let applied_PC =
- mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect 1 realsign)
+ mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect mkRel 1 realsign)
(if dep then [|cstr (2*nrealargs+1) 1|] else [||])) in
let applied_PG =
mkApp (mkVar varP,Array.append (rel_vect 1 nrealargs)
@@ -384,11 +398,11 @@ let build_l2r_rew_scheme dep env ind kind =
(if dep then [|mkRel 2|] else [||])) in
let applied_sym_sym =
mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1;
rel_vect 4 nrealargs;
rel_vect (nrealargs+4) nrealargs;
[|mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1;
rel_vect (nrealargs+4) nrealargs;
rel_vect 4 nrealargs;
[|mkRel 2|]])|]]) in
@@ -411,7 +425,7 @@ let build_l2r_rew_scheme dep env ind kind =
mkApp (eq,[|lift 4 applied_ind;applied_sym_sym;mkRel 1|]),
applied_PR)),
mkApp (sym_involutive,
- Array.append (Context.Rel.to_extended_vect 3 mip.mind_arity_ctxt) [|mkVar varH|]),
+ Array.append (Context.Rel.to_extended_vect mkRel 3 mip.mind_arity_ctxt) [|mkVar varH|]),
[|main_body|])
else
main_body))))))
@@ -450,7 +464,7 @@ let build_l2r_forward_rew_scheme dep env ind kind =
get_sym_eq_data env indu in
let cstr n p =
mkApp (mkConstructUi(indu,1),
- Array.concat [Context.Rel.to_extended_vect n paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -458,12 +472,12 @@ let build_l2r_forward_rew_scheme dep env ind kind =
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_P =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (4*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (4*nrealargs+2) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect (nrealargs+1) nrealargs]) in
let applied_ind_P' =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+1) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+1) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect (2*nrealargs+1) nrealargs]) in
let realsign_P n = lift_rel_context (nrealargs*n+n) realsign in
@@ -541,7 +555,7 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let ((mib,mip as specif),constrargs,realsign,paramsctxt,nrealargs) =
get_non_sym_eq_data env indu in
let cstr n =
- mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in
+ mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
let constrargs_cstr = constrargs@[cstr 0] in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -557,8 +571,8 @@ let build_r2l_forward_rew_scheme dep env ind kind =
applist (mkVar varP,if dep then constrargs_cstr else constrargs) in
let applied_PG =
mkApp (mkVar varP,
- if dep then Context.Rel.to_extended_vect 0 realsign_ind
- else Context.Rel.to_extended_vect 1 realsign) in
+ if dep then Context.Rel.to_extended_vect mkRel 0 realsign_ind
+ else Context.Rel.to_extended_vect mkRel 1 realsign) in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign_ind
@@ -596,7 +610,8 @@ let build_r2l_forward_rew_scheme dep env ind kind =
(**********************************************************************)
let fix_r2l_forward_rew_scheme (c, ctx') =
- let t = Retyping.get_type_of (Global.env()) Evd.empty c in
+ let t = Retyping.get_type_of (Global.env()) Evd.empty (EConstr.of_constr c) in
+ let t = EConstr.Unsafe.to_constr t in
let ctx,_ = decompose_prod_assum t in
match ctx with
| hp :: p :: ind :: indargs ->
@@ -605,9 +620,9 @@ let fix_r2l_forward_rew_scheme (c, ctx') =
(mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 1) p)
(mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 2) hp)
(mkLambda_or_LetIn (RelDecl.map_constr (lift 2) ind)
- (Reductionops.whd_beta Evd.empty
- (applist (c,
- Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))
+ (EConstr.Unsafe.to_constr (Reductionops.whd_beta Evd.empty
+ (EConstr.of_constr (applist (c,
+ Context.Rel.to_extended_list mkRel 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))))
in c', ctx'
| _ -> anomaly (Pp.str "Ill-formed non-dependent left-to-right rewriting scheme")
@@ -762,8 +777,8 @@ let build_congr env (eq,refl,ctx) ind =
(mkNamedLambda varH
(applist
(mkIndU indu,
- Context.Rel.to_extended_list (mip.mind_nrealargs+2) paramsctxt @
- Context.Rel.to_extended_list 0 realsign))
+ Context.Rel.to_extended_list mkRel (mip.mind_nrealargs+2) paramsctxt @
+ Context.Rel.to_extended_list mkRel 0 realsign))
(mkCase (ci,
my_it_mkLambda_or_LetIn_name
(lift_rel_context (mip.mind_nrealargs+3) realsign)
@@ -771,9 +786,9 @@ let build_congr env (eq,refl,ctx) ind =
(Anonymous,
applist
(mkIndU indu,
- Context.Rel.to_extended_list (2*mip.mind_nrealdecls+3)
+ Context.Rel.to_extended_list mkRel (2*mip.mind_nrealdecls+3)
paramsctxt
- @ Context.Rel.to_extended_list 0 realsign),
+ @ Context.Rel.to_extended_list mkRel 0 realsign),
mkApp (eq,
[|mkVar varB;
mkApp (mkVar varf, [|lift (2*mip.mind_nrealdecls+4) b|]);
diff --git a/tactics/equality.ml b/tactics/equality.ml
index d44dcf10d..7ae7446c8 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -6,18 +6,20 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Inductive
open Inductiveops
-open Environ
open Libnames
open Globnames
open Reductionops
@@ -144,7 +146,7 @@ let freeze_initial_evars sigma flags clause =
(* We take evars of the type: this may include old evars! For excluding *)
(* all old evars, including the ones occurring in the rewriting lemma, *)
(* we would have to take the clenv_value *)
- let newevars = Evd.evars_of_term (clenv_type clause) in
+ let newevars = Evarutil.undefined_evars_of_term sigma (clenv_type clause) in
let evars =
fold_undefined (fun evk _ evars ->
if Evar.Set.mem evk newevars then evars
@@ -165,8 +167,9 @@ let side_tac tac sidetac =
let instantiate_lemma_all frzevars gl c ty l l2r concl =
let env = Proofview.Goal.env gl in
+ let sigma = project gl in
let eqclause = pf_apply Clenv.make_clenv_binding gl (c,ty) l in
- let (equiv, args) = decompose_appvect (Clenv.clenv_type eqclause) in
+ let (equiv, args) = decompose_app_vect sigma (Clenv.clenv_type eqclause) in
let arglen = Array.length args in
let () = if arglen < 2 then error "The term provided is not an applied relation." in
let c1 = args.(arglen - 2) in
@@ -304,7 +307,7 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim =
else instantiate_lemma_all frzevars gl c t l l2r concl
in
let typ = match cls with
- | None -> pf_nf_concl gl
+ | None -> pf_concl gl
| Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl)
in
let cs = instantiate_lemma typ in
@@ -326,9 +329,9 @@ let (forward_general_setoid_rewrite_clause, general_setoid_rewrite_clause) = Hoo
let jmeq_same_dom gl = function
| None -> true (* already checked in Hipattern.find_eq_data_decompose *)
| Some t ->
- let rels, t = decompose_prod_assum t in
- let env = Environ.push_rel_context rels (Proofview.Goal.env gl) in
- match decompose_app t with
+ let rels, t = decompose_prod_assum (project gl) t in
+ let env = push_rel_context rels (Proofview.Goal.env gl) in
+ match decompose_app (project gl) t with
| _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2
| _ -> false
@@ -336,6 +339,8 @@ let jmeq_same_dom gl = function
eliminate lbeq on sort_of_gl. *)
let find_elim hdcncl lft2rgt dep cls ot gl =
+ let sigma = project gl in
+ let is_global gr c = Termops.is_global sigma gr c in
let inccl = Option.is_empty cls in
if (is_global Coqlib.glob_eq hdcncl ||
(is_global Coqlib.glob_jmeq hdcncl &&
@@ -343,7 +348,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
|| Flags.version_less_or_equal Flags.V8_2
then
let c =
- match kind_of_term hdcncl with
+ match EConstr.kind sigma hdcncl with
| Ind (ind_sp,u) ->
let pr1 =
lookup_eliminator ind_sp (elimination_sort_of_clause cls gl)
@@ -371,6 +376,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
assert false
in
let Sigma (elim, sigma, p) = Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in
+ let elim = EConstr.of_constr elim in
Sigma ((elim, Safe_typing.empty_private_constants), sigma, p)
else
let scheme_name = match dep, lft2rgt, inccl with
@@ -385,13 +391,14 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
| true, _, true -> rew_r2l_dep_scheme_kind
| true, _, false -> rew_r2l_forward_dep_scheme_kind
in
- match kind_of_term hdcncl with
+ match EConstr.kind sigma hdcncl with
| Ind (ind,u) ->
let c, eff = find_scheme scheme_name ind in
(* MS: cannot use pf_constr_of_global as the eliminator might be generated by side-effect *)
let Sigma (elim, sigma, p) =
Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c)
in
+ let elim = EConstr.of_constr elim in
Sigma ((elim, eff), sigma, p)
| _ -> assert false
@@ -400,12 +407,12 @@ let type_of_clause cls gl = match cls with
| Some id -> pf_get_hyp_typ id gl
let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in
- let isatomic = isProd (whd_zeta evd hdcncl) in
+ let isatomic = isProd evd (whd_zeta evd hdcncl) in
let dep_fun = if isatomic then dependent else dependent_no_evar in
let type_of_cls = type_of_clause cls gl in
- let dep = dep_proof_ok && dep_fun c type_of_cls in
+ let dep = dep_proof_ok && dep_fun evd c type_of_cls in
let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in
let tac =
Proofview.tclEFFECTS effs <*>
@@ -441,8 +448,8 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
let ctype = get_type_of env sigma c in
- let rels, t = decompose_prod_assum (whd_betaiotazeta sigma ctype) in
- match match_with_equality_type t with
+ let rels, t = decompose_prod_assum sigma (whd_betaiotazeta sigma ctype) in
+ match match_with_equality_type sigma t with
| Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *)
let lft2rgt = adjust_rewriting_direction args lft2rgt in
leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels)
@@ -455,9 +462,10 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
end
begin function
| (e, info) ->
+ Proofview.tclEVARMAP >>= fun sigma ->
let env' = push_rel_context rels env in
let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *)
- match match_with_equality_type t' with
+ match match_with_equality_type sigma t' with
| Some (hdcncl,args) ->
let lft2rgt = adjust_rewriting_direction args lft2rgt in
leibniz_rewrite_ebindings_clause cls lft2rgt tac c
@@ -523,7 +531,7 @@ let general_rewrite_clause l2r with_evars ?tac c cl =
let do_hyps =
(* If the term to rewrite uses an hypothesis H, don't rewrite in H *)
let ids gl =
- let ids_in_c = Environ.global_vars_set (Global.env()) (fst c) in
+ let ids_in_c = Termops.global_vars_set (Global.env()) (project gl) (fst c) in
let ids_of_hyps = pf_ids_of_hyps gl in
Id.Set.fold (fun id l -> List.remove Id.equal id l) ids_in_c ids_of_hyps
in
@@ -713,7 +721,6 @@ let _ =
optread = (fun () -> !keep_proof_equalities_for_injection) ;
optwrite = (fun b -> keep_proof_equalities_for_injection := b) }
-
let find_positions env sigma t1 t2 =
let project env sorts posn t1 t2 =
let ty1 = get_type_of env sigma t1 in
@@ -724,7 +731,7 @@ let find_positions env sigma t1 t2 =
let rec findrec sorts posn t1 t2 =
let hd1,args1 = whd_all_stack env sigma t1 in
let hd2,args2 = whd_all_stack env sigma t2 in
- match (kind_of_term hd1, kind_of_term hd2) with
+ match (EConstr.kind sigma hd1, EConstr.kind sigma hd2) with
| Construct ((ind1,i1 as sp1),u1), Construct (sp2,_)
when Int.equal (List.length args1) (constructor_nallargs_env env sp1)
->
@@ -738,8 +745,10 @@ let find_positions env sigma t1 t2 =
let params1,rargs1 = List.chop nparams args1 in
let _,rargs2 = List.chop nparams args2 in
let (mib,mip) = lookup_mind_specif env ind1 in
+ let params1 = List.map EConstr.Unsafe.to_constr params1 in
+ let u1 = EInstance.kind sigma u1 in
let ctxt = (get_constructor ((ind1,u1),mib,mip,params1) i1).cs_args in
- let adjust i = Vars.adjust_rel_to_rel_context ctxt (i+1) - 1 in
+ let adjust i = CVars.adjust_rel_to_rel_context ctxt (i+1) - 1 in
List.flatten
(List.map2_i (fun i -> findrec sorts' ((sp1,adjust i)::posn))
0 rargs1 rargs2)
@@ -850,21 +859,23 @@ let descend_then env sigma head dirn =
let (mib,mip) = lookup_mind_specif env ind in
let cstr = get_constructors env indf in
let dirn_nlams = cstr.(dirn-1).cs_nargs in
- let dirn_env = push_rel_context cstr.(dirn-1).cs_args env in
+ let dirn_env = Environ.push_rel_context cstr.(dirn-1).cs_args env in
(dirn_nlams,
dirn_env,
(fun dirnval (dfltval,resty) ->
- let deparsign = make_arity_signature env true indf in
+ let deparsign = make_arity_signature env sigma true indf in
let p =
it_mkLambda_or_LetIn (lift (mip.mind_nrealargs+1) resty) deparsign in
let build_branch i =
let result = if Int.equal i dirn then dirnval else dfltval in
- it_mkLambda_or_LetIn_name env result cstr.(i-1).cs_args in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.(i-1).cs_args in
+ let args = name_context env sigma cs_args in
+ it_mkLambda_or_LetIn result args in
let brl =
List.map build_branch
(List.interval 1 (Array.length mip.mind_consnames)) in
let ci = make_case_info env ind RegularStyle in
- Inductiveops.make_case_or_project env indf ci p head (Array.of_list brl)))
+ Inductiveops.make_case_or_project env sigma indf ci p head (Array.of_list brl)))
(* Now we need to construct the discriminator, given a discriminable
position. This boils down to:
@@ -899,17 +910,22 @@ let build_selector env sigma dirn c ind special default =
let ind, _ = check_privacy env indp in
let typ = Retyping.get_type_of env sigma default in
let (mib,mip) = lookup_mind_specif env ind in
- let deparsign = make_arity_signature env true indf in
+ let deparsign = make_arity_signature env sigma true indf in
let p = it_mkLambda_or_LetIn typ deparsign in
let cstrs = get_constructors env indf in
let build_branch i =
let endpt = if Int.equal i dirn then special else default in
- it_mkLambda_or_LetIn endpt cstrs.(i-1).cs_args in
+ let args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstrs.(i-1).cs_args in
+ it_mkLambda_or_LetIn endpt args in
let brl =
List.map build_branch(List.interval 1 (Array.length mip.mind_consnames)) in
let ci = make_case_info env ind RegularStyle in
mkCase (ci, p, c, Array.of_list brl)
+let build_coq_False () = EConstr.of_constr (build_coq_False ())
+let build_coq_True () = EConstr.of_constr (build_coq_True ())
+let build_coq_I () = EConstr.of_constr (build_coq_I ())
+
let rec build_discriminator env sigma dirn c = function
| [] ->
let ind = get_type_of env sigma c in
@@ -935,9 +951,9 @@ let rec build_discriminator env sigma dirn c = function
let gen_absurdity id =
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
let hyp_typ = pf_get_hyp_typ id (Proofview.Goal.assume gl) in
- let hyp_typ = pf_nf_evar gl hyp_typ in
- if is_empty_type hyp_typ
+ if is_empty_type sigma hyp_typ
then
simplest_elim (mkVar id)
else
@@ -967,6 +983,7 @@ let discrimination_pf env sigma e (t,t1,t2) discriminator lbeq =
let absurd_term = build_coq_False () in
let eq_elim, eff = ind_scheme_of_eq lbeq in
let sigma, eq_elim = Evd.fresh_global (Global.env ()) sigma eq_elim in
+ let eq_elim = EConstr.of_constr eq_elim in
sigma, (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term),
eff
@@ -976,7 +993,7 @@ let apply_on_clause (f,t) clause =
let sigma = clause.evd in
let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in
let argmv =
- (match kind_of_term (last_arg f_clause.templval.Evd.rebus) with
+ (match EConstr.kind sigma (last_arg f_clause.evd f_clause.templval.Evd.rebus) with
| Meta mv -> mv
| _ -> user_err (str "Ill-formed clause applicator.")) in
clenv_fchain ~with_univs:false argmv f_clause clause
@@ -998,7 +1015,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn =
let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause =
let sigma = eq_clause.evd in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
match find_positions env sigma t1 t2 with
| Inr _ ->
@@ -1008,7 +1025,7 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause =
end }
let onEquality with_evars tac (c,lbindc) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let type_of = pf_unsafe_type_of gl in
let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in
let t = type_of c in
@@ -1023,12 +1040,12 @@ let onEquality with_evars tac (c,lbindc) =
end }
let onNegatedEquality with_evars tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = Tacmach.New.project gl in
let ccl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
- match kind_of_term (hnf_constr env sigma ccl) with
- | Prod (_,t,u) when is_empty_type u ->
+ match EConstr.kind sigma (hnf_constr env sigma ccl) with
+ | Prod (_,t,u) when is_empty_type sigma u ->
tclTHEN introf
(onLastHypId (fun id ->
onEquality with_evars tac (mkVar id,NoBindings)))
@@ -1082,7 +1099,7 @@ let find_sigma_data env s = build_sigma_type ()
*)
let make_tuple env sigma (rterm,rty) lind =
- assert (dependent (mkRel lind) rty);
+ assert (not (noccurn sigma lind rty));
let sigdata = find_sigma_data env (get_sort_of env sigma rty) in
let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in
let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in
@@ -1092,6 +1109,8 @@ let make_tuple env sigma (rterm,rty) lind =
let p = mkLambda (na, a, rty) in
let sigma, exist_term = Evd.fresh_global env sigma sigdata.intro in
let sigma, sig_term = Evd.fresh_global env sigma sigdata.typ in
+ let exist_term = EConstr.of_constr exist_term in
+ let sig_term = EConstr.of_constr sig_term in
sigma,
(applist(exist_term,[a;p;(mkRel lind);rterm]),
applist(sig_term,[a;p]))
@@ -1104,9 +1123,9 @@ let make_tuple env sigma (rterm,rty) lind =
normalization *)
let minimal_free_rels env sigma (c,cty) =
- let cty_rels = free_rels cty in
+ let cty_rels = free_rels sigma cty in
let cty' = simpl env sigma cty in
- let rels' = free_rels cty' in
+ let rels' = free_rels sigma cty' in
if Int.Set.subset cty_rels rels' then
(cty,cty_rels)
else
@@ -1174,19 +1193,18 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt =
error "Cannot solve a unification problem."
else
let (a,p_i_minus_1) = match whd_beta_stack !evdref p_i with
- | (_sigS,[a;p]) -> (a,p)
+ | (_sigS,[a;p]) -> (a, p)
| _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in
let ev = Evarutil.e_new_evar env evdref a in
- let rty = beta_applist(p_i_minus_1,[ev]) in
+ let rty = beta_applist sigma (p_i_minus_1,[ev]) in
let tuple_tail = sigrec_clausal_form (siglen-1) rty in
- match
- Evd.existential_opt_value !evdref
- (destEvar ev)
- with
+ let evopt = match EConstr.kind !evdref ev with Evar _ -> None | _ -> Some ev in
+ match evopt with
| Some w ->
- let w_type = unsafe_type_of env sigma w in
+ let w_type = unsafe_type_of env !evdref w in
if Evarconv.e_cumul env evdref w_type a then
let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in
+ let exist_term = EConstr.of_constr exist_term in
applist(exist_term,[a;p_i_minus_1;w;tuple_tail])
else
error "Cannot solve a unification problem."
@@ -1267,7 +1285,7 @@ let make_iterated_tuple env sigma dflt (z,zty) =
let sigma, (tuple,tuplety) =
List.fold_left (fun (sigma, t) -> make_tuple env sigma t) (sigma, (z,zty)) sorted_rels
in
- assert (closed0 tuplety);
+ assert (closed0 sigma tuplety);
let n = List.length sorted_rels in
let sigma, dfltval = sig_clausal_form env sigma sort_of_zty n tuplety dflt in
sigma, (tuple,tuplety,dfltval)
@@ -1287,56 +1305,46 @@ let build_injector env sigma dflt c cpath =
let sigma, (injcode,resty,_) = build_injrec env sigma dflt c cpath in
sigma, (injcode,resty)
-(*
-let try_delta_expand env sigma t =
- let whdt = whd_all env sigma t in
- let rec hd_rec c =
- match kind_of_term c with
- | Construct _ -> whdt
- | App (f,_) -> hd_rec f
- | Cast (c,_,_) -> hd_rec c
- | _ -> t
- in
- hd_rec whdt
-*)
-
let eq_dec_scheme_kind_name = ref (fun _ -> failwith "eq_dec_scheme undefined")
let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k)
let inject_if_homogenous_dependent_pair ty =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
try
+ let sigma = Tacmach.New.project gl in
let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in
(* fetch the informations of the pair *)
let ceq = Universes.constr_of_global Coqlib.glob_eq in
+ let ceq = EConstr.of_constr ceq in
let sigTconstr () = (Coqlib.build_sigma_type()).Coqlib.typ in
let existTconstr () = (Coqlib.build_sigma_type()).Coqlib.intro in
(* check whether the equality deals with dep pairs or not *)
- let eqTypeDest = fst (decompose_app t) in
- if not (Globnames.is_global (sigTconstr()) eqTypeDest) then raise Exit;
- let hd1,ar1 = decompose_app_vect t1 and
- hd2,ar2 = decompose_app_vect t2 in
- if not (Globnames.is_global (existTconstr()) hd1) then raise Exit;
- if not (Globnames.is_global (existTconstr()) hd2) then raise Exit;
- let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in
+ let eqTypeDest = fst (decompose_app sigma t) in
+ if not (Termops.is_global sigma (sigTconstr()) eqTypeDest) then raise Exit;
+ let hd1,ar1 = decompose_app_vect sigma t1 and
+ hd2,ar2 = decompose_app_vect sigma t2 in
+ if not (Termops.is_global sigma (existTconstr()) hd1) then raise Exit;
+ if not (Termops.is_global sigma (existTconstr()) hd2) then raise Exit;
+ let (ind, _), _ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in
(* check if the user has declared the dec principle *)
(* and compare the fst arguments of the dep pair *)
(* Note: should work even if not an inductive type, but the table only *)
(* knows inductive types *)
- if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) &&
+ if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) ind &&
pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit;
Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"];
let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in
let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing"
["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in
- let c, eff = find_scheme (!eq_dec_scheme_kind_name()) (Univ.out_punivs ind) in
+ let inj2 = EConstr.of_constr inj2 in
+ let c, eff = find_scheme (!eq_dec_scheme_kind_name()) ind in
(* cut with the good equality and prove the requested goal *)
tclTHENLIST
[Proofview.tclEFFECTS eff;
intro;
onLastHyp (fun hyp ->
tclTHENS (cut (mkApp (ceq,new_eq_args)))
- [clear [destVar hyp];
+ [clear [destVar sigma hyp];
Proofview.V82.tactic (Tacmach.refine
(mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))
])]
@@ -1350,7 +1358,7 @@ let inject_if_homogenous_dependent_pair ty =
let simplify_args env sigma t =
(* Quick hack to reduce in arguments of eq only *)
- match decompose_app t with
+ match decompose_app sigma t with
| eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma c1;simpl env sigma c2])
| eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma c1;t2;simpl env sigma c2])
| _ -> t
@@ -1365,6 +1373,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
let sigma, (injbody,resty) = build_injector e_env !evdref t1' (mkVar e) cpath in
let injfun = mkNamedLambda e t injbody in
let sigma,congr = Evd.fresh_global env sigma eq.congr in
+ let congr = EConstr.of_constr congr in
let pf = applist(congr,[t;resty;injfun;t1;t2]) in
let sigma, pf_typ = Typing.type_of env sigma pf in
let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in
@@ -1426,7 +1435,8 @@ let injEq ?(old=false) with_evars clear_flag ipats =
match ipats_style with
| Some ipats ->
Proofview.Goal.enter { enter = begin fun gl ->
- let destopt = match kind_of_term c with
+ let sigma = project gl in
+ let destopt = match EConstr.kind sigma c with
| Var id -> get_previous_hyp_position id gl
| _ -> MoveLast in
let clear_tac =
@@ -1455,7 +1465,7 @@ let injConcl = injClause None false None
let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id)))
let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = clause.evd in
let env = Proofview.Goal.env gl in
match find_positions env sigma t1 t2 with
@@ -1476,9 +1486,9 @@ let dEq with_evars =
dEqThen with_evars (fun clear_flag c x ->
(apply_clear_request clear_flag (use_clear_hyp_by_default ()) c))
-let intro_decomp_eq tac data cl =
+let intro_decomp_eq tac data (c, t) =
Proofview.Goal.enter { enter = begin fun gl ->
- let cl = pf_apply make_clenv_binding gl cl NoBindings in
+ let cl = pf_apply make_clenv_binding gl (c, t) NoBindings in
decompEqThen (fun _ -> tac) data cl
end }
@@ -1516,14 +1526,14 @@ let _ = declare_intro_decomp_eq intro_decomp_eq
*)
-let decomp_tuple_term env c t =
+let decomp_tuple_term env sigma c t =
let rec decomprec inner_code ex exty =
let iterated_decomp =
try
- let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose ex in
+ let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose sigma ex in
let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code])
and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in
- let cdrtyp = beta_applist (p,[car]) in
+ let cdrtyp = beta_applist sigma (p,[car]) in
List.map (fun l -> ((car,a),car_code)::l) (decomprec cdr_code cdr cdrtyp)
with Constr_matching.PatternMatchingFailure ->
[]
@@ -1534,8 +1544,8 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
let sigma = Sigma.to_evar_map sigma in
let typ = get_type_of env sigma dep_pair1 in
(* We find all possible decompositions *)
- let decomps1 = decomp_tuple_term env dep_pair1 typ in
- let decomps2 = decomp_tuple_term env dep_pair2 typ in
+ let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in
+ let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in
(* We adjust to the shortest decomposition *)
let n = min (List.length decomps1) (List.length decomps2) in
let decomp1 = List.nth decomps1 (n-1) in
@@ -1547,10 +1557,10 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
(* We build the expected goal *)
let abst_B =
List.fold_right
- (fun (e,t) body -> lambda_create env (t,subst_term e body)) e1_list b in
- let pred_body = beta_applist(abst_B,proj_list) in
- let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in
- let expected_goal = beta_applist (abst_B,List.map fst e2_list) in
+ (fun (e,t) body -> lambda_create env sigma (t,subst_term sigma e body)) e1_list b in
+ let pred_body = beta_applist sigma (abst_B,proj_list) in
+ let body = mkApp (lambda_create env sigma (typ,pred_body),[|dep_pair1|]) in
+ let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in
(* Simulate now the normalisation treatment made by Logic.mk_refgoals *)
let expected_goal = nf_betaiota sigma expected_goal in
(* Retype to get universes right *)
@@ -1564,7 +1574,7 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
(* on for further iterated sigma-tuples *)
let cutSubstInConcl l2r eqn =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in
@@ -1583,7 +1593,7 @@ let cutSubstInConcl l2r eqn =
end }
let cutSubstInHyp l2r eqn id =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in
@@ -1672,14 +1682,14 @@ exception FoundHyp of (Id.t * constr * bool)
let is_eq_x gl x d =
let id = NamedDecl.get_id d in
try
- let is_var id c = match kind_of_term c with
+ let is_var id c = match EConstr.kind (project gl) c with
| Var id' -> Id.equal id id'
| _ -> false
in
let c = pf_nf_evar gl (NamedDecl.get_type d) in
let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in
- if (is_var x lhs) && not (local_occur_var x rhs) then raise (FoundHyp (id,rhs,true));
- if (is_var x rhs) && not (local_occur_var x lhs) then raise (FoundHyp (id,lhs,false))
+ if (is_var x lhs) && not (local_occur_var (project gl) x rhs) then raise (FoundHyp (id,rhs,true));
+ if (is_var x rhs) && not (local_occur_var (project gl) x lhs) then raise (FoundHyp (id,lhs,false))
with Constr_matching.PatternMatchingFailure ->
()
@@ -1689,6 +1699,7 @@ let is_eq_x gl x d =
let subst_one dep_proof_ok x (hyp,rhs,dir) =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
(* The set of hypotheses using x *)
@@ -1696,7 +1707,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
List.rev (pi3 (List.fold_right (fun dcl (dest,deps,allhyps) ->
let id = NamedDecl.get_id dcl in
if not (Id.equal id hyp)
- && List.exists (fun y -> occur_var_in_decl env y dcl) deps
+ && List.exists (fun y -> occur_var_in_decl env sigma y dcl) deps
then
let id_dest = if !regular_subst_tactic then dest else MoveLast in
(dest,id::deps,(id_dest,id)::allhyps)
@@ -1705,7 +1716,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
hyps
(MoveBefore x,[x],[]))) in (* In practice, no dep hyps before x, so MoveBefore x is good enough *)
(* Decides if x appears in conclusion *)
- let depconcl = occur_var env x concl in
+ let depconcl = occur_var env sigma x concl in
let need_rewrite = not (List.is_empty dephyps) || depconcl in
tclTHENLIST
((if need_rewrite then
@@ -1768,13 +1779,15 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
let find_equations gl =
let gl = Proofview.Goal.assume gl in
let env = Proofview.Goal.env gl in
+ let sigma = project gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let select_equation_name decl =
try
let lbeq,u,(_,x,y) = find_eq_data_decompose (NamedDecl.get_type decl) in
+ let u = EInstance.kind sigma u in
let eq = Universes.constr_of_global_univ (lbeq.eq,u) in
if flags.only_leibniz then restrict_to_eq_and_identity eq;
- match kind_of_term x, kind_of_term y with
+ match EConstr.kind sigma x, EConstr.kind sigma y with
| Var z, _ when not (is_evaluable env (EvalVarRef z)) ->
Some (NamedDecl.get_id decl)
| _, Var z when not (is_evaluable env (EvalVarRef z)) ->
@@ -1791,22 +1804,23 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
let process hyp =
Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.assume gl in
+ let sigma = project gl in
let env = Proofview.Goal.env gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let c = pf_get_hyp hyp gl |> NamedDecl.get_type in
let _,_,(_,x,y) = find_eq_data_decompose c in
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
- if Term.eq_constr x y then Proofview.tclUNIT () else
- match kind_of_term x, kind_of_term y with
- | Var x', _ when not (occur_term x y) && not (is_evaluable env (EvalVarRef x')) ->
+ if EConstr.eq_constr sigma x y then Proofview.tclUNIT () else
+ match EConstr.kind sigma x, EConstr.kind sigma y with
+ | Var x', _ when not (occur_term sigma x y) && not (is_evaluable env (EvalVarRef x')) ->
subst_one flags.rewrite_dependent_proof x' (hyp,y,true)
- | _, Var y' when not (occur_term y x) && not (is_evaluable env (EvalVarRef y')) ->
+ | _, Var y' when not (occur_term sigma y x) && not (is_evaluable env (EvalVarRef y')) ->
subst_one flags.rewrite_dependent_proof y' (hyp,x,false)
| _ ->
Proofview.tclUNIT ()
end }
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let ids = find_equations gl in
tclMAP process ids
end }
@@ -1816,17 +1830,19 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
(* Old implementation, not able to manage configurations like a=b, a=t,
or situations like "a = S b, b = S a", or also accidentally unfolding
let-ins *)
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let test (_,c) =
try
let lbeq,u,(_,x,y) = find_eq_data_decompose c in
+ let u = EInstance.kind sigma u in
let eq = Universes.constr_of_global_univ (lbeq.eq,u) in
if flags.only_leibniz then restrict_to_eq_and_identity eq;
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
- if Term.eq_constr x y then failwith "caught";
- match kind_of_term x with Var x -> x | _ ->
- match kind_of_term y with Var y -> y | _ -> failwith "caught"
+ if EConstr.eq_constr sigma x y then failwith "caught";
+ match EConstr.kind sigma x with Var x -> x | _ ->
+ match EConstr.kind sigma y with Var y -> y | _ -> failwith "caught"
with Constr_matching.PatternMatchingFailure -> failwith "caught" in
let test p = try Some (test p) with Failure _ -> None in
let hyps = pf_hyps_types gl in
@@ -1870,7 +1886,7 @@ let rewrite_assumption_cond cond_eq_term cl =
with | Failure _ | UserError _ -> arec rest gl
end
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.lift gl Sigma.Unsafe.le in
let hyps = Proofview.Goal.hyps gl in
arec hyps gl
diff --git a/tactics/equality.mli b/tactics/equality.mli
index 6a4a8126e..5467b4af2 100644
--- a/tactics/equality.mli
+++ b/tactics/equality.mli
@@ -10,6 +10,7 @@
open Names
open Term
open Evd
+open EConstr
open Environ
open Ind_tables
open Locus
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 59d015fa2..77ed4330c 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -6,12 +6,16 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open Util
open CErrors
open Names
-open Vars
open Term
+open Evd
+open EConstr
+open Vars
open Environ
open Mod_subst
open Globnames
@@ -21,7 +25,6 @@ open Libnames
open Smartlocate
open Misctypes
open Tactypes
-open Evd
open Termops
open Inductiveops
open Typing
@@ -45,23 +48,26 @@ type debug = Debug | Info | Off
exception Bound
-let head_constr_bound t =
- let t = strip_outer_cast t in
- let _,ccl = decompose_prod_assum t in
- let hd,args = decompose_app ccl in
- match kind_of_term hd with
- | Const _ | Ind _ | Construct _ | Var _ -> hd
- | Proj (p, _) -> mkConst (Projection.constant p)
- | _ -> raise Bound
-
-let head_constr c =
- try head_constr_bound c with Bound -> error "Bound head variable."
-
-let decompose_app_bound t =
- let t = strip_outer_cast t in
- let _,ccl = decompose_prod_assum t in
- let hd,args = decompose_app_vect ccl in
- match kind_of_term hd with
+let head_constr_bound sigma t =
+ let t = strip_outer_cast sigma t in
+ let _,ccl = decompose_prod_assum sigma t in
+ let hd,args = decompose_app sigma ccl in
+ match EConstr.kind sigma hd with
+ | Const (c, _) -> ConstRef c
+ | Ind (i, _) -> IndRef i
+ | Construct (c, _) -> ConstructRef c
+ | Var id -> VarRef id
+ | Proj (p, _) -> ConstRef (Projection.constant p)
+ | _ -> raise Bound
+
+let head_constr sigma c =
+ try head_constr_bound sigma c with Bound -> error "Bound head variable."
+
+let decompose_app_bound sigma t =
+ let t = strip_outer_cast sigma t in
+ let _,ccl = decompose_prod_assum sigma t in
+ let hd,args = decompose_app_vect sigma ccl in
+ match EConstr.kind sigma hd with
| Const (c,u) -> ConstRef c, args
| Ind (i,u) -> IndRef i, args
| Construct (c,u) -> ConstructRef c, args
@@ -255,8 +261,8 @@ let rebuild_dn st se =
in
{ se with sentry_bnet = dn' }
-let lookup_tacs concl st se =
- let l' = Bounded_net.lookup st se.sentry_bnet concl in
+let lookup_tacs sigma concl st se =
+ let l' = Bounded_net.lookup sigma st se.sentry_bnet concl in
let sl' = List.stable_sort pri_order_int l' in
List.merge pri_order_int se.sentry_nopat sl'
@@ -267,10 +273,10 @@ let is_transparent_gr (ids, csts) = function
| ConstRef cst -> Cpred.mem cst csts
| IndRef _ | ConstructRef _ -> false
-let strip_params env c =
- match kind_of_term c with
+let strip_params env sigma c =
+ match EConstr.kind sigma c with
| App (f, args) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Const (p,_) ->
let cb = lookup_constant p env in
(match cb.Declarations.const_proj with
@@ -289,7 +295,7 @@ let instantiate_hint env sigma p =
let sigma = Evd.merge_context_set univ_flexible sigma ctx in
let cl = mk_clenv_from_env env sigma None (c,cty) in
{cl with templval =
- { cl.templval with rebus = strip_params env cl.templval.rebus };
+ { cl.templval with rebus = strip_params env sigma cl.templval.rebus };
env = empty_env}
in
let code = match p.code.obj with
@@ -473,11 +479,11 @@ val empty : ?name:hint_db_name -> transparent_state -> bool -> t
val find : global_reference -> t -> search_entry
val map_none : secvars:Id.Pred.t -> t -> full_hint list
val map_all : secvars:Id.Pred.t -> global_reference -> t -> full_hint list
-val map_existential : secvars:Id.Pred.t ->
+val map_existential : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
-val map_eauto : secvars:Id.Pred.t ->
+val map_eauto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
-val map_auto : secvars:Id.Pred.t ->
+val map_auto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
val add_one : env -> evar_map -> hint_entry -> t -> t
val add_list : env -> evar_map -> hint_entry list -> t -> t
@@ -533,21 +539,32 @@ struct
(** Warn about no longer typable hint? *)
None
- let match_mode m arg =
+ let head_evar sigma c =
+ let rec hrec c = match EConstr.kind sigma c with
+ | Evar (evk,_) -> evk
+ | Case (_,_,c,_) -> hrec c
+ | App (c,_) -> hrec c
+ | Cast (c,_,_) -> hrec c
+ | Proj (p, c) -> hrec c
+ | _ -> raise Evarutil.NoHeadEvar
+ in
+ hrec c
+
+ let match_mode sigma m arg =
match m with
- | ModeInput -> not (occur_existential arg)
+ | ModeInput -> not (occur_existential sigma arg)
| ModeNoHeadEvar ->
- Evarutil.(try ignore(head_evar arg); false
- with NoHeadEvar -> true)
+ (try ignore(head_evar sigma arg); false
+ with Evarutil.NoHeadEvar -> true)
| ModeOutput -> true
- let matches_mode args mode =
+ let matches_mode sigma args mode =
Array.length mode == Array.length args &&
- Array.for_all2 match_mode mode args
+ Array.for_all2 (match_mode sigma) mode args
- let matches_modes args modes =
+ let matches_modes sigma args modes =
if List.is_empty modes then true
- else List.exists (matches_mode args) modes
+ else List.exists (matches_mode sigma args) modes
let merge_entry secvars db nopat pat =
let h = List.sort pri_order_int (List.map snd db.hintdb_nopat) in
@@ -563,24 +580,24 @@ struct
merge_entry secvars db se.sentry_nopat se.sentry_pat
(** Precondition: concl has no existentials *)
- let map_auto ~secvars (k,args) concl db =
+ let map_auto sigma ~secvars (k,args) concl db =
let se = find k db in
let st = if db.use_dn then (Some db.hintdb_state) else None in
- let pat = lookup_tacs concl st se in
+ let pat = lookup_tacs sigma concl st se in
merge_entry secvars db [] pat
- let map_existential ~secvars (k,args) concl db =
+ let map_existential sigma ~secvars (k,args) concl db =
let se = find k db in
- if matches_modes args se.sentry_mode then
+ if matches_modes sigma args se.sentry_mode then
merge_entry secvars db se.sentry_nopat se.sentry_pat
else merge_entry secvars db [] []
(* [c] contains an existential *)
- let map_eauto ~secvars (k,args) concl db =
+ let map_eauto sigma ~secvars (k,args) concl db =
let se = find k db in
- if matches_modes args se.sentry_mode then
+ if matches_modes sigma args se.sentry_mode then
let st = if db.use_dn then Some db.hintdb_state else None in
- let pat = lookup_tacs concl st se in
+ let pat = lookup_tacs sigma concl st se in
merge_entry secvars db [] pat
else merge_entry secvars db [] []
@@ -745,8 +762,8 @@ let _ = Summary.declare_summary "search"
(* Auxiliary functions to prepare AUTOHINT objects *)
(**************************************************************************)
-let rec nb_hyp c = match kind_of_term c with
- | Prod(_,_,c2) -> if noccurn 1 c2 then 1+(nb_hyp c2) else nb_hyp c2
+let rec nb_hyp sigma c = match EConstr.kind sigma c with
+ | Prod(_,_,c2) -> if noccurn sigma 1 c2 then 1+(nb_hyp sigma c2) else nb_hyp sigma c2
| _ -> 0
(* adding and removing tactics in the search table *)
@@ -763,19 +780,19 @@ let secvars_of_idset s =
Id.Pred.add id p
else p) s Id.Pred.empty
-let secvars_of_constr env c =
- secvars_of_idset (global_vars_set env c)
+let secvars_of_constr env sigma c =
+ secvars_of_idset (Termops.global_vars_set env sigma c)
let secvars_of_global env gr =
secvars_of_idset (vars_of_global_reference env gr)
let make_exact_entry env sigma info poly ?(name=PathAny) (c, cty, ctx) =
- let secvars = secvars_of_constr env c in
- let cty = strip_outer_cast cty in
- match kind_of_term cty with
+ let secvars = secvars_of_constr env sigma c in
+ let cty = strip_outer_cast sigma cty in
+ match EConstr.kind sigma cty with
| Prod _ -> failwith "make_exact_entry"
| _ ->
- let pat = Patternops.pattern_of_constr env sigma cty in
+ let pat = Patternops.pattern_of_constr env sigma (EConstr.to_constr sigma cty) in
let hd =
try head_pattern_bound pat
with BoundPattern -> failwith "make_exact_entry"
@@ -792,18 +809,18 @@ let make_exact_entry env sigma info poly ?(name=PathAny) (c, cty, ctx) =
let make_apply_entry env sigma (eapply,hnf,verbose) info poly ?(name=PathAny) (c, cty, ctx) =
let cty = if hnf then hnf_constr env sigma cty else cty in
- match kind_of_term cty with
+ match EConstr.kind sigma cty with
| Prod _ ->
let sigma' = Evd.merge_context_set univ_flexible sigma ctx in
let ce = mk_clenv_from_env env sigma' None (c,cty) in
let c' = clenv_type (* ~reduce:false *) ce in
- let pat = Patternops.pattern_of_constr env ce.evd c' in
+ let pat = Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma c') in
let hd =
try head_pattern_bound pat
with BoundPattern -> failwith "make_apply_entry" in
let nmiss = List.length (clenv_missing ce) in
- let secvars = secvars_of_constr env c in
- let pri = match info.hint_priority with None -> nb_hyp cty + nmiss | Some p -> p in
+ let secvars = secvars_of_constr env sigma c in
+ let pri = match info.hint_priority with None -> nb_hyp sigma' cty + nmiss | Some p -> p in
let pat = match info.hint_pattern with
| Some p -> snd p | None -> pat
in
@@ -816,7 +833,7 @@ let make_apply_entry env sigma (eapply,hnf,verbose) info poly ?(name=PathAny) (c
else begin
if not eapply then failwith "make_apply_entry";
if verbose then
- Feedback.msg_info (str "the hint: eapply " ++ pr_lconstr c ++
+ Feedback.msg_info (str "the hint: eapply " ++ pr_leconstr_env env sigma' c ++
str " will only be used by eauto");
(Some hd,
{ pri; poly; pat = Some pat; name;
@@ -833,7 +850,7 @@ let pr_hint_term env sigma ctx = function
| IsGlobRef gr -> pr_global gr
| IsConstr (c, ctx) ->
let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in
- pr_constr_env env sigma c
+ pr_econstr_env env sigma c
(** We need an object to record the side-effect of registering
global universes associated with a hint. *)
@@ -859,7 +876,8 @@ let fresh_global_or_constr env sigma poly cr =
let isgr, (c, ctx) =
match cr with
| IsGlobRef gr ->
- true, Universes.fresh_global_instance env gr
+ let (c, ctx) = Universes.fresh_global_instance env gr in
+ true, (EConstr.of_constr c, ctx)
| IsConstr (c, ctx) -> false, (c, ctx)
in
if poly then (c, ctx)
@@ -882,7 +900,7 @@ let make_resolves env sigma flags info poly ?name cr =
in
if List.is_empty ents then
user_err ~hdr:"Hint"
- (pr_lconstr c ++ spc() ++
+ (pr_leconstr_env env sigma c ++ spc() ++
(if pi1 flags then str"cannot be used as a hint."
else str "can be used as a hint only for eauto."));
ents
@@ -924,6 +942,7 @@ let make_extern pri pat tacast =
code = with_uid (Extern tacast) })
let make_mode ref m =
+ let open Term in
let ty = Global.type_of_global_unsafe ref in
let ctx, t = decompose_prod ty in
let n = List.length ctx in
@@ -938,14 +957,14 @@ let make_trivial env sigma poly ?(name=PathAny) r =
let c,ctx = fresh_global_or_constr env sigma poly r in
let sigma = Evd.merge_context_set univ_flexible sigma ctx in
let t = hnf_constr env sigma (unsafe_type_of env sigma c) in
- let hd = head_of_constr_reference (head_constr t) in
+ let hd = head_constr sigma t in
let ce = mk_clenv_from_env env sigma None (c,t) in
(Some hd, { pri=1;
poly = poly;
- pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce));
+ pat = Some (Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma (clenv_type ce)));
name = name;
db = None;
- secvars = secvars_of_constr env c;
+ secvars = secvars_of_constr env sigma c;
code= with_uid (Res_pf_THEN_trivial_fail(c,t,ctx)) })
@@ -1039,14 +1058,16 @@ let cache_autohint (kn, obj) =
let subst_autohint (subst, obj) =
let subst_key gr =
let (lab'', elab') = subst_global subst gr in
+ let elab' = EConstr.of_constr elab' in
let gr' =
- (try head_of_constr_reference (head_constr_bound elab')
+ (try head_constr_bound Evd.empty elab'
with Bound -> lab'')
in if gr' == gr then gr else gr'
in
let subst_hint (k,data as hint) =
let k' = Option.smartmap subst_key k in
let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ let subst_mps subst c = EConstr.of_constr (subst_mps subst (EConstr.Unsafe.to_constr c)) in
let code' = match data.code.obj with
| Res_pf (c,t,ctx) ->
let c' = subst_mps subst c in
@@ -1218,29 +1239,32 @@ let prepare_hint check (poly,local) env init (sigma,c) =
It is actually a bit stupid to generalize over evars since the first
thing make_resolves will do is to re-instantiate the products *)
let sigma, subst = Evd.nf_univ_variables sigma in
- let c = Vars.subst_univs_constr subst (Evarutil.nf_evar sigma c) in
- let c = drop_extra_implicit_args c in
- let vars = ref (collect_vars c) in
+ let c = Evarutil.nf_evar sigma c in
+ let c = EConstr.Unsafe.to_constr c in
+ let c = CVars.subst_univs_constr subst c in
+ let c = EConstr.of_constr c in
+ let c = drop_extra_implicit_args sigma c in
+ let vars = ref (collect_vars sigma c) in
let subst = ref [] in
- let rec find_next_evar c = match kind_of_term c with
+ let rec find_next_evar c = match EConstr.kind sigma c with
| Evar (evk,args as ev) ->
(* We skip the test whether args is the identity or not *)
- let t = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let t = List.fold_right (fun (e,id) c -> replace_term e id c) !subst t in
- if not (closed0 c) then
+ let t = existential_type sigma ev in
+ let t = List.fold_right (fun (e,id) c -> replace_term sigma e id c) !subst t in
+ if not (closed0 sigma c) then
error "Hints with holes dependent on a bound variable not supported.";
- if occur_existential t then
+ if occur_existential sigma t then
(* Not clever enough to construct dependency graph of evars *)
error "Not clever enough to deal with evars dependent in other evars.";
raise (Found (c,t))
- | _ -> Constr.iter find_next_evar c in
+ | _ -> EConstr.iter sigma find_next_evar c in
let rec iter c =
try find_next_evar c; c
with Found (evar,t) ->
let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in
vars := Id.Set.add id !vars;
subst := (evar,mkVar id)::!subst;
- mkNamedLambda id t (iter (replace_term evar (mkVar id) c)) in
+ mkNamedLambda id t (iter (replace_term sigma evar (mkVar id) c)) in
let c' = iter c in
if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c';
let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in
@@ -1321,7 +1345,7 @@ let add_hints local dbnames0 h =
let expand_constructor_hints env sigma lems =
List.map_append (fun (evd,lem) ->
- match kind_of_term lem with
+ match EConstr.kind sigma lem with
| Ind (ind,u) ->
List.init (nconstructors ind)
(fun i ->
@@ -1351,7 +1375,7 @@ let make_local_hint_db env sigma ts eapply lems =
(Sigma.to_evar_map sigma, c)
in
let lems = List.map map lems in
- let sign = Environ.named_context env in
+ let sign = EConstr.named_context env in
let ts = match ts with
| None -> Hint_db.transparent_state (searchtable_map "core")
| Some ts -> ts
@@ -1376,7 +1400,7 @@ let make_db_list dbnames =
(* Functions for printing the hints *)
(**************************************************************************)
-let pr_hint_elt (c, _, _) = pr_constr c
+let pr_hint_elt (c, _, _) = pr_econstr c
let pr_hint h = match h.obj with
| Res_pf (c, _) -> (str"simple apply " ++ pr_hint_elt c)
@@ -1426,15 +1450,15 @@ let pr_hint_ref ref = pr_hint_list_for_head ref
(* Print all hints associated to head id in any database *)
-let pr_hint_term cl =
+let pr_hint_term sigma cl =
try
let dbs = current_db () in
let valid_dbs =
let fn = try
- let hdc = decompose_app_bound cl in
- if occur_existential cl then
- Hint_db.map_existential ~secvars:Id.Pred.full hdc cl
- else Hint_db.map_auto ~secvars:Id.Pred.full hdc cl
+ let hdc = decompose_app_bound sigma cl in
+ if occur_existential sigma cl then
+ Hint_db.map_existential sigma ~secvars:Id.Pred.full hdc cl
+ else Hint_db.map_auto sigma ~secvars:Id.Pred.full hdc cl
with Bound -> Hint_db.map_none ~secvars:Id.Pred.full
in
let fn db = List.map (fun x -> 0, x) (fn db) in
@@ -1455,7 +1479,7 @@ let pr_applicable_hint () =
match glss.Evd.it with
| [] -> CErrors.error "No focused goal."
| g::_ ->
- pr_hint_term (Goal.V82.concl glss.Evd.sigma g)
+ pr_hint_term glss.Evd.sigma (Goal.V82.concl glss.Evd.sigma g)
let pp_hint_mode = function
| ModeInput -> str"+"
diff --git a/tactics/hints.mli b/tactics/hints.mli
index 05d41adfe..467fd46d5 100644
--- a/tactics/hints.mli
+++ b/tactics/hints.mli
@@ -10,6 +10,7 @@ open Pp
open Util
open Names
open Term
+open EConstr
open Environ
open Globnames
open Decl_kinds
@@ -24,11 +25,11 @@ open Vernacexpr
exception Bound
-val decompose_app_bound : constr -> global_reference * constr array
+val decompose_app_bound : evar_map -> constr -> global_reference * constr array
type debug = Debug | Info | Off
-val secvars_of_hyps : Context.Named.t -> Id.Pred.t
+val secvars_of_hyps : ('c, 't) Context.Named.pt -> Id.Pred.t
val empty_hint_info : 'a hint_info_gen
@@ -110,16 +111,16 @@ module Hint_db :
(** All hints associated to the reference, respecting modes if evars appear in the
arguments, _not_ using the discrimination net. *)
- val map_existential : secvars:Id.Pred.t ->
+ val map_existential : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
(** All hints associated to the reference, respecting modes if evars appear in the
arguments and using the discrimination net. *)
- val map_eauto : secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list
+ val map_eauto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list
(** All hints associated to the reference, respecting modes if evars appear in the
arguments. *)
- val map_auto : secvars:Id.Pred.t ->
+ val map_auto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
val add_one : env -> evar_map -> hint_entry -> t -> t
@@ -182,7 +183,7 @@ val add_hints : locality_flag -> hint_db_name list -> hints_entry -> unit
val prepare_hint : bool (* Check no remaining evars *) ->
(bool * bool) (* polymorphic or monomorphic, local or global *) ->
- env -> evar_map -> open_constr -> hint_term
+ env -> evar_map -> evar_map * constr -> hint_term
(** [make_exact_entry info (c, ctyp, ctx)].
[c] is the term given as an exact proof to solve the goal;
@@ -231,7 +232,7 @@ val make_resolves :
If the hyp cannot be used as a Hint, the empty list is returned. *)
val make_resolve_hyp :
- env -> evar_map -> Context.Named.Declaration.t -> hint_entry list
+ env -> evar_map -> named_declaration -> hint_entry list
(** [make_extern pri pattern tactic_expr] *)
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index 27af7200b..851554b83 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -12,6 +12,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Inductiveops
open Constr_matching
open Coqlib
@@ -31,9 +32,9 @@ module RelDecl = Context.Rel.Declaration
-- Eduardo (6/8/97). *)
-type 'a matching_function = constr -> 'a option
+type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option
-type testing_function = constr -> bool
+type testing_function = Evd.evar_map -> EConstr.constr -> bool
let mkmeta n = Nameops.make_ident "X" (Some n)
let meta1 = mkmeta 1
@@ -43,11 +44,11 @@ let meta4 = mkmeta 4
let op2bool = function Some _ -> true | None -> false
-let match_with_non_recursive_type t =
- match kind_of_term t with
+let match_with_non_recursive_type sigma t =
+ match EConstr.kind sigma t with
| App _ ->
- let (hdapp,args) = decompose_app t in
- (match kind_of_term hdapp with
+ let (hdapp,args) = decompose_app sigma t in
+ (match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
if (Global.lookup_mind (fst ind)).mind_finite == Decl_kinds.CoFinite then
Some (hdapp,args)
@@ -56,21 +57,21 @@ let match_with_non_recursive_type t =
| _ -> None)
| _ -> None
-let is_non_recursive_type t = op2bool (match_with_non_recursive_type t)
+let is_non_recursive_type sigma t = op2bool (match_with_non_recursive_type sigma t)
(* Test dependencies *)
(* NB: we consider also the let-in case in the following function,
since they may appear in types of inductive constructors (see #2629) *)
-let rec has_nodep_prod_after n c =
- match kind_of_term c with
+let rec has_nodep_prod_after n sigma c =
+ match EConstr.kind sigma c with
| Prod (_,_,b) | LetIn (_,_,_,b) ->
- ( n>0 || not (dependent (mkRel 1) b))
- && (has_nodep_prod_after (n-1) b)
+ ( n>0 || Vars.noccurn sigma 1 b)
+ && (has_nodep_prod_after (n-1) sigma b)
| _ -> true
-let has_nodep_prod = has_nodep_prod_after 0
+let has_nodep_prod sigma c = has_nodep_prod_after 0 sigma c
(* A general conjunctive type is a non-recursive with-no-indices inductive
type with only one constructor and no dependencies between argument;
@@ -87,9 +88,11 @@ let is_lax_conjunction = function
| Some false -> true
| _ -> false
-let match_with_one_constructor style onlybinary allow_rec t =
- let (hdapp,args) = decompose_app t in
- let res = match kind_of_term hdapp with
+let prod_assum sigma t = fst (decompose_prod_assum sigma t)
+
+let match_with_one_constructor sigma style onlybinary allow_rec t =
+ let (hdapp,args) = decompose_app sigma t in
+ let res = match EConstr.kind sigma hdapp with
| Ind ind ->
let (mib,mip) = Global.lookup_inductive (fst ind) in
if Int.equal (Array.length mip.mind_consnames) 1
@@ -98,21 +101,21 @@ let match_with_one_constructor style onlybinary allow_rec t =
then
if is_strict_conjunction style (* strict conjunction *) then
let ctx =
- (prod_assum (snd
- (decompose_prod_n_assum mib.mind_nparams mip.mind_nf_lc.(0)))) in
+ (prod_assum sigma (snd
+ (decompose_prod_n_assum sigma mib.mind_nparams (EConstr.of_constr mip.mind_nf_lc.(0))))) in
if
List.for_all
(fun decl -> let c = RelDecl.get_type decl in
is_local_assum decl &&
- isRel c &&
- Int.equal (destRel c) mib.mind_nparams) ctx
+ isRel sigma c &&
+ Int.equal (destRel sigma c) mib.mind_nparams) ctx
then
Some (hdapp,args)
else None
else
- let ctyp = prod_applist mip.mind_nf_lc.(0) args in
- let cargs = List.map RelDecl.get_type (prod_assum ctyp) in
- if not (is_lax_conjunction style) || has_nodep_prod ctyp then
+ let ctyp = Termops.prod_applist sigma (EConstr.of_constr mip.mind_nf_lc.(0)) args in
+ let cargs = List.map RelDecl.get_type (prod_assum sigma ctyp) in
+ if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then
(* Record or non strict conjunction *)
Some (hdapp,List.rev cargs)
else
@@ -125,28 +128,29 @@ let match_with_one_constructor style onlybinary allow_rec t =
| Some (hdapp, [_; _]) -> res
| _ -> None
-let match_with_conjunction ?(strict=false) ?(onlybinary=false) t =
- match_with_one_constructor (Some strict) onlybinary false t
+let match_with_conjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ match_with_one_constructor sigma (Some strict) onlybinary false t
-let match_with_record t =
- match_with_one_constructor None false false t
+let match_with_record sigma t =
+ match_with_one_constructor sigma None false false t
-let is_conjunction ?(strict=false) ?(onlybinary=false) t =
- op2bool (match_with_conjunction ~strict ~onlybinary t)
+let is_conjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ op2bool (match_with_conjunction sigma ~strict ~onlybinary t)
-let is_record t =
- op2bool (match_with_record t)
+let is_record sigma t =
+ op2bool (match_with_record sigma t)
-let match_with_tuple t =
- let t = match_with_one_constructor None false true t in
+let match_with_tuple sigma t =
+ let t = match_with_one_constructor sigma None false true t in
Option.map (fun (hd,l) ->
- let ind = destInd hd in
+ let ind = destInd sigma hd in
+ let ind = on_snd (fun u -> EInstance.kind sigma u) ind in
let (mib,mip) = Global.lookup_pinductive ind in
let isrec = mis_is_recursive (fst ind,mib,mip) in
(hd,l,isrec)) t
-let is_tuple t =
- op2bool (match_with_tuple t)
+let is_tuple sigma t =
+ op2bool (match_with_tuple sigma t)
(* A general disjunction type is a non-recursive with-no-indices inductive
type with of which all constructors have a single argument;
@@ -154,14 +158,15 @@ let is_tuple t =
"Inductive I A1 ... An := C1 (_:A1) | ... | Cn : (_:An)" *)
let test_strict_disjunction n lc =
+ let open Term in
Array.for_all_i (fun i c ->
match (prod_assum (snd (decompose_prod_n_assum n c))) with
| [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i)
| _ -> false) 0 lc
-let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
- let (hdapp,args) = decompose_app t in
- let res = match kind_of_term hdapp with
+let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ let res = match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
let car = constructors_nrealargs ind in
let (mib,mip) = Global.lookup_inductive ind in
@@ -176,7 +181,7 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
None
else
let cargs =
- Array.map (fun ar -> pi2 (destProd (prod_applist ar args)))
+ Array.map (fun ar -> pi2 (destProd sigma (prod_applist sigma (EConstr.of_constr ar) args)))
mip.mind_nf_lc in
Some (hdapp,Array.to_list cargs)
else
@@ -187,48 +192,48 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
| Some (hdapp,[_; _]) -> res
| _ -> None
-let is_disjunction ?(strict=false) ?(onlybinary=false) t =
- op2bool (match_with_disjunction ~strict ~onlybinary t)
+let is_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ op2bool (match_with_disjunction ~strict ~onlybinary sigma t)
(* An empty type is an inductive type, possible with indices, that has no
constructors *)
-let match_with_empty_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_empty_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
let nconstr = Array.length mip.mind_consnames in
if Int.equal nconstr 0 then Some hdapp else None
| _ -> None
-let is_empty_type t = op2bool (match_with_empty_type t)
+let is_empty_type sigma t = op2bool (match_with_empty_type sigma t)
(* This filters inductive types with one constructor with no arguments;
Parameters and indices are allowed *)
-let match_with_unit_or_eq_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_unit_or_eq_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind , _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
let constr_types = mip.mind_nf_lc in
let nconstr = Array.length mip.mind_consnames in
- let zero_args c = Int.equal (nb_prod c) mib.mind_nparams in
+ let zero_args c = Int.equal (nb_prod sigma (EConstr.of_constr c)) mib.mind_nparams in
if Int.equal nconstr 1 && zero_args constr_types.(0) then
Some hdapp
else
None
| _ -> None
-let is_unit_or_eq_type t = op2bool (match_with_unit_or_eq_type t)
+let is_unit_or_eq_type sigma t = op2bool (match_with_unit_or_eq_type sigma t)
(* A unit type is an inductive type with no indices but possibly
(useless) parameters, and that has no arguments in its unique
constructor *)
-let is_unit_type t =
- match match_with_conjunction t with
+let is_unit_type sigma t =
+ match match_with_conjunction sigma t with
| Some (_,[]) -> true
| _ -> false
@@ -276,13 +281,13 @@ let coq_refl_jm_pattern =
open Globnames
-let is_matching x y = is_matching (Global.env ()) Evd.empty x y
-let matches x y = matches (Global.env ()) Evd.empty x y
+let is_matching sigma x y = is_matching (Global.env ()) sigma x y
+let matches sigma x y = matches (Global.env ()) sigma x y
-let match_with_equation t =
- if not (isApp t) then raise NoEquationFound;
- let (hdapp,args) = destApp t in
- match kind_of_term hdapp with
+let match_with_equation sigma t =
+ if not (isApp sigma t) then raise NoEquationFound;
+ let (hdapp,args) = destApp sigma t in
+ match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
if eq_gr (IndRef ind) glob_eq then
Some (build_coq_eq_data()),hdapp,
@@ -298,11 +303,11 @@ let match_with_equation t =
let constr_types = mip.mind_nf_lc in
let nconstr = Array.length mip.mind_consnames in
if Int.equal nconstr 1 then
- if is_matching coq_refl_leibniz1_pattern constr_types.(0) then
+ if is_matching sigma coq_refl_leibniz1_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, MonomorphicLeibnizEq(args.(0),args.(1))
- else if is_matching coq_refl_leibniz2_pattern constr_types.(0) then
+ else if is_matching sigma coq_refl_leibniz2_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, PolymorphicLeibnizEq(args.(0),args.(1),args.(2))
- else if is_matching coq_refl_jm_pattern constr_types.(0) then
+ else if is_matching sigma coq_refl_jm_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, HeterogenousEq(args.(0),args.(1),args.(2),args.(3))
else raise NoEquationFound
else raise NoEquationFound
@@ -318,57 +323,57 @@ let is_inductive_equality ind =
let nconstr = Array.length mip.mind_consnames in
Int.equal nconstr 1 && Int.equal (constructor_nrealargs (ind,1)) 0
-let match_with_equality_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
+let match_with_equality_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
| Ind (ind,_) when is_inductive_equality ind -> Some (hdapp,args)
| _ -> None
-let is_equality_type t = op2bool (match_with_equality_type t)
+let is_equality_type sigma t = op2bool (match_with_equality_type sigma t)
(* Arrows/Implication/Negation *)
(** X1 -> X2 **)
let coq_arrow_pattern = mkPattern (mkGArrow (mkGPatVar "X1") (mkGPatVar "X2"))
-let match_arrow_pattern t =
- let result = matches coq_arrow_pattern t in
+let match_arrow_pattern sigma t =
+ let result = matches sigma coq_arrow_pattern t in
match Id.Map.bindings result with
| [(m1,arg);(m2,mind)] ->
assert (Id.equal m1 meta1 && Id.equal m2 meta2); (arg, mind)
| _ -> anomaly (Pp.str "Incorrect pattern matching")
-let match_with_imp_term c=
- match kind_of_term c with
- | Prod (_,a,b) when not (dependent (mkRel 1) b) ->Some (a,b)
+let match_with_imp_term sigma c =
+ match EConstr.kind sigma c with
+ | Prod (_,a,b) when Vars.noccurn sigma 1 b -> Some (a,b)
| _ -> None
-let is_imp_term c = op2bool (match_with_imp_term c)
+let is_imp_term sigma c = op2bool (match_with_imp_term sigma c)
-let match_with_nottype t =
+let match_with_nottype sigma t =
try
- let (arg,mind) = match_arrow_pattern t in
- if is_empty_type mind then Some (mind,arg) else None
+ let (arg,mind) = match_arrow_pattern sigma t in
+ if is_empty_type sigma mind then Some (mind,arg) else None
with PatternMatchingFailure -> None
-let is_nottype t = op2bool (match_with_nottype t)
+let is_nottype sigma t = op2bool (match_with_nottype sigma t)
(* Forall *)
-let match_with_forall_term c=
- match kind_of_term c with
+let match_with_forall_term sigma c=
+ match EConstr.kind sigma c with
| Prod (nam,a,b) -> Some (nam,a,b)
| _ -> None
-let is_forall_term c = op2bool (match_with_forall_term c)
+let is_forall_term sigma c = op2bool (match_with_forall_term sigma c)
-let match_with_nodep_ind t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_nodep_ind sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
if Array.length (mib.mind_packets)>1 then None else
- let nodep_constr = has_nodep_prod_after mib.mind_nparams in
+ let nodep_constr c = has_nodep_prod_after mib.mind_nparams sigma (EConstr.of_constr c) in
if Array.for_all nodep_constr mip.mind_nf_lc then
let params=
if Int.equal mip.mind_nrealargs 0 then args else
@@ -378,24 +383,24 @@ let match_with_nodep_ind t =
None
| _ -> None
-let is_nodep_ind t=op2bool (match_with_nodep_ind t)
+let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t)
-let match_with_sigma_type t=
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_sigma_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
if Int.equal (Array.length (mib.mind_packets)) 1 &&
(Int.equal mip.mind_nrealargs 0) &&
(Int.equal (Array.length mip.mind_consnames)1) &&
- has_nodep_prod_after (mib.mind_nparams+1) mip.mind_nf_lc.(0) then
+ has_nodep_prod_after (mib.mind_nparams+1) sigma (EConstr.of_constr mip.mind_nf_lc.(0)) then
(*allowing only 1 existential*)
Some (hdapp,args)
else
None
| _ -> None
-let is_sigma_type t=op2bool (match_with_sigma_type t)
+let is_sigma_type sigma t = op2bool (match_with_sigma_type sigma t)
(***** Destructing patterns bound to some theory *)
@@ -408,17 +413,17 @@ let rec first_match matcher = function
(*** Equality *)
-let match_eq eqn (ref, hetero) =
+let match_eq sigma eqn (ref, hetero) =
let ref =
try Lazy.force ref
with e when CErrors.noncritical e -> raise PatternMatchingFailure
in
- match kind_of_term eqn with
+ match EConstr.kind sigma eqn with
| App (c, [|t; x; y|]) ->
- if not hetero && is_global ref c then PolymorphicLeibnizEq (t, x, y)
+ if not hetero && Termops.is_global sigma ref c then PolymorphicLeibnizEq (t, x, y)
else raise PatternMatchingFailure
| App (c, [|t; x; t'; x'|]) ->
- if hetero && is_global ref c then HeterogenousEq (t, x, t', x')
+ if hetero && Termops.is_global sigma ref c then HeterogenousEq (t, x, t', x')
else raise PatternMatchingFailure
| _ -> raise PatternMatchingFailure
@@ -430,9 +435,9 @@ let equalities =
(coq_jmeq_ref, true), check_jmeq_loaded, build_coq_jmeq_data;
(coq_identity_ref, false), no_check, build_coq_identity_data]
-let find_eq_data eqn = (* fails with PatternMatchingFailure *)
- let d,k = first_match (match_eq eqn) equalities in
- let hd,u = destInd (fst (destApp eqn)) in
+let find_eq_data sigma eqn = (* fails with PatternMatchingFailure *)
+ let d,k = first_match (match_eq sigma eqn) equalities in
+ let hd,u = destInd sigma (fst (destApp sigma eqn)) in
d,u,k
let extract_eq_args gl = function
@@ -444,13 +449,13 @@ let extract_eq_args gl = function
else raise PatternMatchingFailure
let find_eq_data_decompose gl eqn =
- let (lbeq,u,eq_args) = find_eq_data eqn in
+ let (lbeq,u,eq_args) = find_eq_data (project gl) eqn in
(lbeq,u,extract_eq_args gl eq_args)
let find_this_eq_data_decompose gl eqn =
let (lbeq,u,eq_args) =
try (*first_match (match_eq eqn) inversible_equalities*)
- find_eq_data eqn
+ find_eq_data (project gl) eqn
with PatternMatchingFailure ->
user_err (str "No primitive equality found.") in
let eq_args =
@@ -477,12 +482,12 @@ let dest_nf_eq gls eqn =
(*** Sigma-types *)
-let match_sigma ex =
- match kind_of_term ex with
- | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_exist_ref) f ->
- build_sigma (), (snd (destConstruct f), a, p, car, cdr)
- | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_existT_ref) f ->
- build_sigma_type (), (snd (destConstruct f), a, p, car, cdr)
+let match_sigma sigma ex =
+ match EConstr.kind sigma ex with
+ | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_exist_ref) f ->
+ build_sigma (), (snd (destConstruct sigma f), a, p, car, cdr)
+ | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_existT_ref) f ->
+ build_sigma_type (), (snd (destConstruct sigma f), a, p, car, cdr)
| _ -> raise PatternMatchingFailure
let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *)
@@ -492,12 +497,12 @@ let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *)
let coq_sig_pattern =
lazy (mkPattern (mkGAppRef coq_sig_ref [mkGPatVar "X1"; mkGPatVar "X2"]))
-let match_sigma t =
- match Id.Map.bindings (matches (Lazy.force coq_sig_pattern) t) with
+let match_sigma sigma t =
+ match Id.Map.bindings (matches sigma (Lazy.force coq_sig_pattern) t) with
| [(_,a); (_,p)] -> (a,p)
| _ -> anomaly (Pp.str "Unexpected pattern")
-let is_matching_sigma t = is_matching (Lazy.force coq_sig_pattern) t
+let is_matching_sigma sigma t = is_matching sigma (Lazy.force coq_sig_pattern) t
(*** Decidable equalities *)
@@ -529,26 +534,26 @@ let coq_eqdec_rev_pattern = coq_eqdec ~sum:coq_or_ref ~rev:true
let op_or = coq_or_ref
let op_sum = coq_sumbool_ref
-let match_eqdec t =
+let match_eqdec sigma t =
let eqonleft,op,subst =
- try true,op_sum,matches (Lazy.force coq_eqdec_inf_pattern) t
+ try true,op_sum,matches sigma (Lazy.force coq_eqdec_inf_pattern) t
with PatternMatchingFailure ->
- try false,op_sum,matches (Lazy.force coq_eqdec_inf_rev_pattern) t
+ try false,op_sum,matches sigma (Lazy.force coq_eqdec_inf_rev_pattern) t
with PatternMatchingFailure ->
- try true,op_or,matches (Lazy.force coq_eqdec_pattern) t
+ try true,op_or,matches sigma (Lazy.force coq_eqdec_pattern) t
with PatternMatchingFailure ->
- false,op_or,matches (Lazy.force coq_eqdec_rev_pattern) t in
+ false,op_or,matches sigma (Lazy.force coq_eqdec_rev_pattern) t in
match Id.Map.bindings subst with
| [(_,typ);(_,c1);(_,c2)] ->
- eqonleft, Universes.constr_of_global (Lazy.force op), c1, c2, typ
+ eqonleft, EConstr.of_constr (Universes.constr_of_global (Lazy.force op)), c1, c2, typ
| _ -> anomaly (Pp.str "Unexpected pattern")
(* Patterns "~ ?" and "? -> False" *)
let coq_not_pattern = lazy (mkPattern (mkGAppRef coq_not_ref [mkGHole]))
let coq_imp_False_pattern = lazy (mkPattern (mkGArrow mkGHole (mkGRef coq_False_ref)))
-let is_matching_not t = is_matching (Lazy.force coq_not_pattern) t
-let is_matching_imp_False t = is_matching (Lazy.force coq_imp_False_pattern) t
+let is_matching_not sigma t = is_matching sigma (Lazy.force coq_not_pattern) t
+let is_matching_imp_False sigma t = is_matching sigma (Lazy.force coq_imp_False_pattern) t
(* Remark: patterns that have references to the standard library must
be evaluated lazily (i.e. at the time they are used, not a the time
diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli
index 7cc41f1b9..dd09c3a4d 100644
--- a/tactics/hipattern.mli
+++ b/tactics/hipattern.mli
@@ -8,6 +8,8 @@
open Names
open Term
+open Evd
+open EConstr
open Coqlib
(** High-order patterns *)
@@ -40,8 +42,8 @@ open Coqlib
also work on ad-hoc disjunctions introduced by the user.
(Eduardo, 6/8/97). *)
-type 'a matching_function = constr -> 'a option
-type testing_function = constr -> bool
+type 'a matching_function = evar_map -> constr -> 'a option
+type testing_function = evar_map -> constr -> bool
val match_with_non_recursive_type : (constr * constr list) matching_function
val is_non_recursive_type : testing_function
@@ -113,39 +115,39 @@ type equation_kind =
exception NoEquationFound
val match_with_equation:
- constr -> coq_eq_data option * constr * equation_kind
+ evar_map -> constr -> coq_eq_data option * constr * equation_kind
(***** Destructing patterns bound to some theory *)
(** Match terms [eq A t u], [identity A t u] or [JMeq A t A u]
Returns associated lemmas and [A,t,u] or fails PatternMatchingFailure *)
-val find_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr ->
- coq_eq_data * Univ.universe_instance * (types * constr * constr)
+val find_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr ->
+ coq_eq_data * EInstance.t * (types * constr * constr)
(** Idem but fails with an error message instead of PatternMatchingFailure *)
-val find_this_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr ->
- coq_eq_data * Univ.universe_instance * (types * constr * constr)
+val find_this_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr ->
+ coq_eq_data * EInstance.t * (types * constr * constr)
(** A variant that returns more informative structure on the equality found *)
-val find_eq_data : constr -> coq_eq_data * Univ.universe_instance * equation_kind
+val find_eq_data : evar_map -> constr -> coq_eq_data * EInstance.t * equation_kind
(** Match a term of the form [(existT A P t p)]
Returns associated lemmas and [A,P,t,p] *)
-val find_sigma_data_decompose : constr ->
- coq_sigma_data * (Univ.universe_instance * constr * constr * constr * constr)
+val find_sigma_data_decompose : evar_map -> constr ->
+ coq_sigma_data * (EInstance.t * constr * constr * constr * constr)
(** Match a term of the form [{x:A|P}], returns [A] and [P] *)
-val match_sigma : constr -> constr * constr
+val match_sigma : evar_map -> constr -> constr * constr
-val is_matching_sigma : constr -> bool
+val is_matching_sigma : evar_map -> constr -> bool
(** Match a decidable equality judgement (e.g [{t=u:>T}+{~t=u}]), returns
[t,u,T] and a boolean telling if equality is on the left side *)
-val match_eqdec : constr -> bool * constr * constr * constr * constr
+val match_eqdec : evar_map -> constr -> bool * constr * constr * constr * constr
(** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *)
-val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr)
+val dest_nf_eq : ('a, 'r) Proofview.Goal.t -> constr -> (constr * constr * constr)
(** Match a negation *)
-val is_matching_not : constr -> bool
-val is_matching_imp_False : constr -> bool
+val is_matching_not : evar_map -> constr -> bool
+val is_matching_imp_False : evar_map -> constr -> bool
diff --git a/tactics/inv.ml b/tactics/inv.ml
index e7d8249e4..904a17417 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -12,10 +12,11 @@ open Util
open Names
open Nameops
open Term
-open Vars
open Termops
-open Namegen
open Environ
+open EConstr
+open Vars
+open Namegen
open Inductiveops
open Printer
open Retyping
@@ -32,8 +33,9 @@ module NamedDecl = Context.Named.Declaration
let var_occurs_in_pf gl id =
let env = Proofview.Goal.env gl in
- occur_var env id (Proofview.Goal.concl gl) ||
- List.exists (occur_var_in_decl env id) (Proofview.Goal.hyps gl)
+ let sigma = project gl in
+ occur_var env sigma id (Proofview.Goal.concl gl) ||
+ List.exists (occur_var_in_decl env sigma id) (Proofview.Goal.hyps gl)
(* [make_inv_predicate (ity,args) C]
@@ -73,9 +75,10 @@ let make_inv_predicate env evd indf realargs id status concl =
| NoDep ->
(* We push the arity and leave concl unchanged *)
let hyps_arity,_ = get_arity env indf in
+ let hyps_arity = List.map (fun d -> map_rel_decl EConstr.of_constr d) hyps_arity in
(hyps_arity,concl)
| Dep dflt_concl ->
- if not (occur_var env id concl) then
+ if not (occur_var env !evd id concl) then
user_err ~hdr:"make_inv_predicate"
(str "Current goal does not depend on " ++ pr_id id ++ str".");
(* We abstract the conclusion of goal with respect to
@@ -87,11 +90,11 @@ let make_inv_predicate env evd indf realargs id status concl =
| None ->
let sort = get_sort_family_of env !evd concl in
let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in
- let p = make_arity env true indf sort in
+ let p = make_arity env !evd true indf sort in
let evd',(p,ptyp) = Unification.abstract_list_all env
!evd p concl (realargs@[mkVar id])
in evd := evd'; p in
- let hyps,bodypred = decompose_lam_n_assum (nrealargs+1) pred in
+ let hyps,bodypred = decompose_lam_n_assum !evd (nrealargs+1) pred in
(* We lift to make room for the equations *)
(hyps,lift nrealargs bodypred)
in
@@ -109,7 +112,7 @@ let make_inv_predicate env evd indf realargs id status concl =
let ai = lift nhyps ai in
let (xi, ti) = compute_eqn env' !evd nhyps n ai in
let (lhs,eqnty,rhs) =
- if closed0 ti then
+ if closed0 !evd ti then
(xi,ti,ai)
else
let sigma, res = make_iterated_tuple env' !evd ai (xi,ti) in
@@ -117,17 +120,19 @@ let make_inv_predicate env evd indf realargs id status concl =
in
let eq_term = eqdata.Coqlib.eq in
let eq = Evarutil.evd_comb1 (Evd.fresh_global env) evd eq_term in
+ let eq = EConstr.of_constr eq in
let eqn = applist (eq,[eqnty;lhs;rhs]) in
let eqns = (Anonymous, lift n eqn) :: eqns in
let refl_term = eqdata.Coqlib.refl in
let refl_term = Evarutil.evd_comb1 (Evd.fresh_global env) evd refl_term in
+ let refl_term = EConstr.of_constr refl_term in
let refl = mkApp (refl_term, [|eqnty; rhs|]) in
let _ = Evarutil.evd_comb1 (Typing.type_of env) evd refl in
let args = refl :: args in
build_concl eqns args (succ n) restlist
in
let (newconcl, args) = build_concl [] [] 0 realargs in
- let predicate = it_mkLambda_or_LetIn_name env newconcl hyps in
+ let predicate = it_mkLambda_or_LetIn newconcl (name_context env !evd hyps) in
let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in
(* OK - this predicate should now be usable by res_elimination_then to
do elimination on the conclusion. *)
@@ -183,7 +188,7 @@ let dependent_hyps env id idlist gl =
| d::l ->
(* Update the type of id1: it may have been subject to rewriting *)
let d = pf_get_hyp (NamedDecl.get_id d) gl in
- if occur_var_in_decl env id d
+ if occur_var_in_decl env (project gl) id d
then d :: dep_rec l
else dep_rec l
in
@@ -268,7 +273,7 @@ Nota: with Inversion_clear, only four useless hypotheses
let generalizeRewriteIntros as_mode tac depids id =
Proofview.tclENV >>= fun env ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let dids = dependent_hyps env id depids gl in
let reintros = if as_mode then intros_replacing else intros_possibly_replacing in
(tclTHENLIST
@@ -283,7 +288,7 @@ let error_too_many_names pats =
tclZEROMSG ~loc (
str "Unexpected " ++
str (String.plural (List.length pats) "introduction pattern") ++
- str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (fst (run_delayed env Evd.empty c)))) pats ++
+ str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed env Evd.empty c))))) pats ++
str ".")
let get_names (allow_conj,issimple) (loc, pat as x) = match pat with
@@ -333,15 +338,16 @@ let remember_first_eq id x = if !x == MoveLast then x := MoveAfter id
let projectAndApply as_mode thin avoid id eqname names depids =
let subst_hyp l2r id =
- tclTHEN (tclTRY(rewriteInConcl l2r (mkVar id)))
+ tclTHEN (tclTRY(rewriteInConcl l2r (EConstr.mkVar id)))
(if thin then clear [id] else (remember_first_eq id eqname; tclIDTAC))
in
let substHypIfVariable tac id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
(** We only look at the type of hypothesis "id" *)
let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in
let (t,t1,t2) = Hipattern.dest_nf_eq gl hyp in
- match (kind_of_term t1, kind_of_term t2) with
+ match (EConstr.kind sigma t1, EConstr.kind sigma t2) with
| Var id1, _ -> generalizeRewriteIntros as_mode (subst_hyp true id) depids id1
| _, Var id2 -> generalizeRewriteIntros as_mode (subst_hyp false id) depids id2
| _ -> tac id
@@ -368,11 +374,11 @@ let projectAndApply as_mode thin avoid id eqname names depids =
(* and apply a trailer which again try to substitute *)
(fun id ->
dEqThen false (deq_trailer id)
- (Some (None,ElimOnConstr (mkVar id,NoBindings))))
+ (Some (None,ElimOnConstr (EConstr.mkVar id,NoBindings))))
id
let nLastDecls i tac =
- Proofview.Goal.nf_enter { enter = begin fun gl -> tac (nLastDecls gl i) end }
+ Proofview.Goal.enter { enter = begin fun gl -> tac (nLastDecls gl i) end }
(* Introduction of the equations on arguments
othin: discriminates Simple Inversion, Inversion and Inversion_clear
@@ -380,7 +386,7 @@ let nLastDecls i tac =
Some thin: the equations are rewritten, and cleared if thin is true *)
let rewrite_equations as_mode othin neqns names ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let (depids,nodepids) = split_dep_and_nodep ba.Tacticals.assums gl in
let first_eq = ref MoveLast in
let avoid = if as_mode then List.map NamedDecl.get_id nodepids else [] in
@@ -430,7 +436,7 @@ let rewrite_equations_tac as_mode othin id neqns names ba =
tac
let raw_inversion inv_kind id status names =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let sigma = Sigma.to_evar_map sigma in
let env = Proofview.Goal.env gl in
@@ -448,11 +454,11 @@ let raw_inversion inv_kind id status names =
make_inv_predicate env evdref indf realargs id status concl in
let sigma = !evdref in
let (cut_concl,case_tac) =
- if status != NoDep && (dependent c concl) then
- Reduction.beta_appvect elim_predicate (Array.of_list (realargs@[c])),
+ if status != NoDep && (dependent sigma c concl) then
+ Reductionops.beta_applist sigma (elim_predicate, realargs@[c]),
case_then_using
else
- Reduction.beta_appvect elim_predicate (Array.of_list realargs),
+ Reductionops.beta_applist sigma (elim_predicate, realargs),
case_nodep_then_using
in
let refined id =
@@ -467,7 +473,7 @@ let raw_inversion inv_kind id status names =
[case_tac names
(introCaseAssumsThen false (* ApplyOn not supported by inversion *)
(rewrite_equations_tac as_mode inv_kind id neqns))
- (Some elim_predicate) ind (c, t);
+ (Some elim_predicate) ind (c,t);
onLastHypId (fun id -> tclTHEN (refined id) reflexivity)])
in
Sigma.Unsafe.of_pair (tac, sigma)
@@ -511,15 +517,17 @@ let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id)
* back to their places in the hyp-list. *)
let invIn k names ids id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let hyps = List.map (fun id -> pf_get_hyp id gl) ids in
let concl = Proofview.Goal.concl gl in
- let nb_prod_init = nb_prod concl in
+ let sigma = project gl in
+ let nb_prod_init = nb_prod sigma concl in
let intros_replace_ids =
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
+ let concl = pf_concl gl in
+ let sigma = project gl in
let nb_of_new_hyp =
- nb_prod concl - (List.length hyps + nb_prod_init)
+ nb_prod sigma concl - (List.length hyps + nb_prod_init)
in
if nb_of_new_hyp < 1 then
intros_replacing ids
diff --git a/tactics/inv.mli b/tactics/inv.mli
index df629e7c9..446a62f6d 100644
--- a/tactics/inv.mli
+++ b/tactics/inv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Misctypes
open Tactypes
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 10fc5076c..daa962f1d 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -11,15 +11,16 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Evd
open Printer
open Reductionops
open Entries
open Inductiveops
-open Environ
open Tacmach.New
open Clenv
open Declare
@@ -33,7 +34,7 @@ module NamedDecl = Context.Named.Declaration
let no_inductive_inconstr env sigma constr =
(str "Cannot recognize an inductive predicate in " ++
- pr_lconstr_env env sigma constr ++
+ pr_leconstr_env env sigma constr ++
str "." ++ spc () ++ str "If there is one, may be the structure of the arity" ++
spc () ++ str "or of the type of constructors" ++ spc () ++
str "is hidden by constant definitions.")
@@ -116,13 +117,13 @@ let max_prefix_sign lid sign =
| id::l -> snd (max_rec (id, sign_prefix id sign) l)
*)
let rec add_prods_sign env sigma t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (na,c1,b) ->
- let id = id_of_name_using_hdchar env t na in
+ let id = id_of_name_using_hdchar env sigma t na in
let b'= subst1 (mkVar id) b in
add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b'
| LetIn (na,c1,t1,b) ->
- let id = id_of_name_using_hdchar env t na in
+ let id = id_of_name_using_hdchar env sigma t na in
let b'= subst1 (mkVar id) b in
add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b'
| _ -> (env,t)
@@ -146,7 +147,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
let p = next_ident_away (Id.of_string "P") allvars in
let pty,goal =
if dep_option then
- let pty = make_arity env true indf sort in
+ let pty = make_arity env sigma true indf sort in
let goal =
mkProd
(Anonymous, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1]))
@@ -154,10 +155,11 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
pty,goal
else
let i = mkAppliedInd ind in
- let ivars = global_vars env i in
+ let ivars = global_vars env sigma i in
let revargs,ownsign =
fold_named_context
(fun env d (revargs,hyps) ->
+ let d = map_named_decl EConstr.of_constr d in
let id = NamedDecl.get_id d in
if Id.List.mem id ivars then
((mkVar id)::revargs, Context.Named.add d hyps)
@@ -192,7 +194,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
in
assert
(List.subset
- (global_vars env invGoal)
+ (global_vars env sigma invGoal)
(ids_of_named_context (named_context invEnv)));
(*
user_err ~hdr:"lemma_inversion"
@@ -208,6 +210,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let ownSign = ref begin
fold_named_context
(fun env d sign ->
+ let d = map_named_decl EConstr.of_constr d in
if mem_named_context_val (NamedDecl.get_id d) global_named_context then sign
else Context.Named.add d sign)
invEnv ~init:Context.Named.empty
@@ -216,18 +219,19 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let { sigma=sigma } = Proof.V82.subgoals pf in
let sigma = Evd.nf_constraints sigma in
let rec fill_holes c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Evar (e,args) ->
let h = next_ident_away (Id.of_string "H") !avoid in
let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in
avoid := h::!avoid;
ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign;
applist (mkVar h, inst)
- | _ -> Constr.map fill_holes c
+ | _ -> EConstr.map sigma fill_holes c
in
let c = fill_holes pfterm in
(* warning: side-effect on ownSign *)
let invProof = it_mkNamedLambda_or_LetIn c !ownSign in
+ let invProof = EConstr.Unsafe.to_constr invProof in
let p = Evarutil.nf_evars_universes sigma invProof in
p, Evd.universe_context sigma
@@ -258,26 +262,28 @@ let add_inversion_lemma_exn na com comsort bool tac =
let lemInv id c gls =
try
- let clause = mk_clenv_type_of gls c in
- let clause = clenv_constrain_last_binding (mkVar id) clause in
+ let open Tacmach in
+ let clause = mk_clenv_from_env (pf_env gls) (project gls) None (c, pf_unsafe_type_of gls c) in
+ let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in
Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls
with
| NoSuchBinding ->
user_err
- (hov 0 (pr_constr c ++ spc () ++ str "does not refer to an inversion lemma."))
+ (hov 0 (pr_econstr_env (Refiner.pf_env gls) (Refiner.project gls) c ++ spc () ++ str "does not refer to an inversion lemma."))
| UserError (a,b) ->
user_err ~hdr:"LemInv"
(str "Cannot refine current goal with the lemma " ++
- pr_lconstr_env (Refiner.pf_env gls) (Refiner.project gls) c)
+ pr_leconstr_env (Refiner.pf_env gls) (Refiner.project gls) c)
let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id
let lemInvIn id c ids =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let hyps = List.map (fun id -> pf_get_hyp id gl) ids in
let intros_replace_ids =
let concl = Proofview.Goal.concl gl in
- let nb_of_new_hyp = nb_prod concl - List.length ids in
+ let sigma = project gl in
+ let nb_of_new_hyp = nb_prod sigma concl - List.length ids in
if nb_of_new_hyp < 1 then
intros_replacing ids
else
diff --git a/tactics/leminv.mli b/tactics/leminv.mli
index c6ed9606f..26d4ac994 100644
--- a/tactics/leminv.mli
+++ b/tactics/leminv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Constrexpr
open Misctypes
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index c5562b326..90b7d6581 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -11,6 +11,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Termops
open Declarations
open Tacmach
@@ -83,7 +84,7 @@ let nLastDecls n gl =
try List.firstn n (pf_hyps gl)
with Failure _ -> error "Not enough hypotheses in the goal."
-let nLastHypsId n gl = List.map NamedDecl.get_id (nLastDecls n gl)
+let nLastHypsId n gl = List.map (NamedDecl.get_id) (nLastDecls n gl)
let nLastHyps n gl = List.map mkVar (nLastHypsId n gl)
let onNthDecl m tac gl = tac (nthDecl m gl) gl
@@ -157,7 +158,7 @@ type branch_args = {
type branch_assumptions = {
ba : branch_args; (* the branch args *)
- assums : Context.Named.t} (* the list of assumptions introduced *)
+ assums : named_context} (* the list of assumptions introduced *)
open Misctypes
@@ -226,6 +227,7 @@ 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
| Prod (_,_,c), recarg::rest ->
@@ -263,40 +265,7 @@ let pf_with_evars glsev k gls =
tclTHEN (Refiner.tclEVARS evd) (k a) gls
let pf_constr_of_global gr k =
- pf_with_evars (fun gls -> pf_apply Evd.fresh_global gls gr) k
-
-(* computing the case/elim combinators *)
-
-let gl_make_elim ind gl =
- let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in
- pf_apply Evd.fresh_global gl gr
-
-let gl_make_case_dep ind gl =
- let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in
- let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true
- (elimination_sort_of_goal gl)
- in
- (Sigma.to_evar_map sigma, r)
-
-let gl_make_case_nodep ind gl =
- let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in
- let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false
- (elimination_sort_of_goal gl)
- in
- (Sigma.to_evar_map sigma, r)
-
-let make_elim_branch_assumptions ba gl =
- let assums =
- try List.rev (List.firstn ba.nassums (pf_hyps gl))
- with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in
- { ba = ba; assums = assums }
-
-let elim_on_ba tac ba gl = tac (make_elim_branch_assumptions ba gl) gl
-
-let make_case_branch_assumptions = make_elim_branch_assumptions
-
-let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl
-
+ pf_with_evars (fun gls -> on_snd EConstr.of_constr (pf_apply Evd.fresh_global gls gr)) k
(** Tacticals of Ltac defined directly in term of Proofview *)
module New = struct
@@ -544,7 +513,7 @@ module New = struct
Proofview.Unsafe.tclEVARS sigma <*> tac >>= check_evars_if
let tclDELAYEDWITHHOLES check x tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let Sigma (x, sigma, _) = x.delayed env sigma in
@@ -588,13 +557,13 @@ module New = struct
let onLastHyp = onNthHyp 1
let onNthDecl m tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Proofview.tclUNIT (nthDecl m gl) >>= tac
end }
let onLastDecl = onNthDecl 1
let ifOnHyp pred tac1 tac2 id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let typ = Tacmach.New.pf_get_hyp_typ id gl in
if pred (id,typ) then
tac1 id
@@ -602,7 +571,7 @@ module New = struct
tac2 id
end }
- let onHyps find tac = Proofview.Goal.nf_enter { enter = begin fun gl -> tac (find.enter gl) end }
+ let onHyps find tac = Proofview.Goal.enter { enter = begin fun gl -> tac (find.enter gl) end }
let afterHyp id tac =
Proofview.Goal.enter { enter = begin fun gl ->
@@ -635,25 +604,26 @@ module New = struct
(* c should be of type A1->.. An->B with B an inductive definition *)
let general_elim_then_using mk_elim
isrec allnames tac predicate ind (c, t) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma, elim = (mk_elim ind).enter gl in
+ let ind = on_snd (fun u -> EInstance.kind sigma u) ind in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
- (Proofview.Goal.nf_enter { enter = begin fun gl ->
- let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in
+ (Proofview.Goal.enter { enter = begin fun gl ->
+ let indclause = mk_clenv_from gl (c, t) in
(* applying elimination_scheme just a little modified *)
- let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in
+ let elimclause = mk_clenv_from gl (elim,Tacmach.New.pf_unsafe_type_of gl elim) in
let indmv =
- match kind_of_term (last_arg elimclause.templval.Evd.rebus) with
+ match EConstr.kind elimclause.evd (last_arg elimclause.evd elimclause.templval.Evd.rebus) with
| Meta mv -> mv
| _ -> anomaly (str"elimination")
in
let pmv =
- let p, _ = decompose_app elimclause.templtyp.Evd.rebus in
- match kind_of_term p with
+ let p, _ = decompose_app elimclause.evd elimclause.templtyp.Evd.rebus in
+ match EConstr.kind elimclause.evd p with
| Meta p -> p
| _ ->
let name_elim =
- match kind_of_term elim with
+ match EConstr.kind sigma elim with
| Const (kn, _) -> string_of_con kn
| Var id -> string_of_id id
| _ -> "\b"
@@ -670,9 +640,9 @@ module New = struct
| None -> elimclause'
| Some p -> clenv_unify ~flags Reduction.CONV (mkMeta pmv) p elimclause'
in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in
+ let clenv' = clenv_unique_resolver ~flags elimclause' gl in
let after_tac i =
- let (hd,largs) = decompose_app clenv'.templtyp.Evd.rebus in
+ let (hd,largs) = decompose_app clenv'.evd clenv'.templtyp.Evd.rebus in
let ba = { branchsign = branchsigns.(i);
branchnames = brnames.(i);
nassums = List.length branchsigns.(i);
@@ -689,8 +659,66 @@ module New = struct
(Proofview.tclEXTEND [] tclIDTAC branchtacs)
end }) end }
+ let elimination_sort_of_goal gl =
+ (** Retyping will expand evars anyway. *)
+ let c = Proofview.Goal.concl (Goal.assume gl) in
+ pf_apply Retyping.get_sort_family_of gl c
+
+ let elimination_sort_of_hyp id gl =
+ (** Retyping will expand evars anyway. *)
+ let c = pf_get_hyp_typ id (Goal.assume gl) in
+ pf_apply Retyping.get_sort_family_of gl c
+
+ let elimination_sort_of_clause id gl = match id with
+ | None -> elimination_sort_of_goal gl
+ | Some id -> elimination_sort_of_hyp id gl
+
+ (* computing the case/elim combinators *)
+
+ let gl_make_elim ind = { enter = begin fun gl ->
+ let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in
+ let (sigma, c) = pf_apply Evd.fresh_global gl gr in
+ (sigma, EConstr.of_constr c)
+ end }
+
+ let gl_make_case_dep (ind, u) = { enter = begin fun gl ->
+ let sigma = Sigma.Unsafe.of_evar_map (project gl) in
+ let u = EInstance.kind (project gl) u in
+ let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) true
+ (elimination_sort_of_goal gl)
+ in
+ (Sigma.to_evar_map sigma, EConstr.of_constr r)
+ end }
+
+ let gl_make_case_nodep (ind, u) = { enter = begin fun gl ->
+ let sigma = Sigma.Unsafe.of_evar_map (project gl) in
+ let u = EInstance.kind (project gl) u in
+ let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) false
+ (elimination_sort_of_goal gl)
+ in
+ (Sigma.to_evar_map sigma, EConstr.of_constr r)
+ end }
+
+ let make_elim_branch_assumptions ba hyps =
+ let assums =
+ try List.rev (List.firstn ba.nassums hyps)
+ with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in
+ { ba = ba; assums = assums }
+
+ let elim_on_ba tac ba =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in
+ tac branches
+ end }
+
+ let case_on_ba tac ba =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in
+ tac branches
+ end }
+
let elimination_then tac c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let (ind,t) = pf_reduce_to_quantified_ind gl (pf_unsafe_type_of gl c) in
let isrec,mkelim =
match (Global.lookup_mind (fst (fst ind))).mind_record with
@@ -706,37 +734,12 @@ module New = struct
let case_nodep_then_using =
general_elim_then_using gl_make_case_nodep false
- let elim_on_ba tac ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let branches = Tacmach.New.of_old (make_elim_branch_assumptions ba) gl in
- tac branches
- end }
-
- let case_on_ba tac ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let branches = Tacmach.New.of_old (make_case_branch_assumptions ba) gl in
- tac branches
- end }
-
- let elimination_sort_of_goal gl =
- (** Retyping will expand evars anyway. *)
- let c = Proofview.Goal.concl (Goal.assume gl) in
- pf_apply Retyping.get_sort_family_of gl c
-
- let elimination_sort_of_hyp id gl =
- (** Retyping will expand evars anyway. *)
- let c = pf_get_hyp_typ id (Goal.assume gl) in
- pf_apply Retyping.get_sort_family_of gl c
-
- let elimination_sort_of_clause id gl = match id with
- | None -> elimination_sort_of_goal gl
- | Some id -> elimination_sort_of_hyp id gl
-
let pf_constr_of_global ref tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let (sigma, c) = Evd.fresh_global env sigma ref in
+ let c = EConstr.of_constr c in
Proofview.Unsafe.tclEVARS sigma <*> (tac c)
end }
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 7aacc52f3..3b90ec514 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -9,6 +9,7 @@
open Pp
open Names
open Term
+open EConstr
open Tacmach
open Proof_type
open Locus
@@ -59,29 +60,29 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic
val onNthHypId : int -> (Id.t -> tactic) -> tactic
val onNthHyp : int -> (constr -> tactic) -> tactic
-val onNthDecl : int -> (Context.Named.Declaration.t -> tactic) -> tactic
+val onNthDecl : int -> (named_declaration -> tactic) -> tactic
val onLastHypId : (Id.t -> tactic) -> tactic
val onLastHyp : (constr -> tactic) -> tactic
-val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic
+val onLastDecl : (named_declaration -> tactic) -> tactic
val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic
val onNLastHyps : int -> (constr list -> tactic) -> tactic
-val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic
+val onNLastDecls : int -> (named_context -> tactic) -> tactic
val lastHypId : goal sigma -> Id.t
val lastHyp : goal sigma -> constr
-val lastDecl : goal sigma -> Context.Named.Declaration.t
+val lastDecl : goal sigma -> named_declaration
val nLastHypsId : int -> goal sigma -> Id.t list
val nLastHyps : int -> goal sigma -> constr list
-val nLastDecls : int -> goal sigma -> Context.Named.t
+val nLastDecls : int -> goal sigma -> named_context
-val afterHyp : Id.t -> goal sigma -> Context.Named.t
+val afterHyp : Id.t -> goal sigma -> named_context
val ifOnHyp : (Id.t * types -> bool) ->
(Id.t -> tactic) -> (Id.t -> tactic) ->
Id.t -> tactic
-val onHyps : (goal sigma -> Context.Named.t) ->
- (Context.Named.t -> tactic) -> tactic
+val onHyps : (goal sigma -> named_context) ->
+ (named_context -> tactic) -> tactic
(** {6 Tacticals applying to goal components } *)
@@ -97,7 +98,7 @@ val onClauseLR : (Id.t option -> tactic) -> clause -> tactic
(** {6 Elimination tacticals. } *)
-type branch_args = {
+type branch_args = private {
ity : pinductive; (** the type we were eliminating on *)
largs : constr list; (** its arguments *)
branchnum : int; (** the branch number *)
@@ -107,9 +108,9 @@ type branch_args = {
true=assumption, false=let-in *)
branchnames : intro_patterns}
-type branch_assumptions = {
+type branch_assumptions = private {
ba : branch_args; (** the branch args *)
- assums : Context.Named.t} (** the list of assumptions introduced *)
+ assums : named_context} (** the list of assumptions introduced *)
(** [get_and_check_or_and_pattern loc pats branchsign] returns an appropriate
error message if |pats| <> |branchsign|; extends them if no pattern is given
@@ -123,7 +124,7 @@ val fix_empty_or_and_pattern : int ->
delayed_open_constr or_and_intro_pattern_expr ->
delayed_open_constr or_and_intro_pattern_expr
-val compute_constructor_signatures : rec_flag -> pinductive -> bool list array
+val compute_constructor_signatures : rec_flag -> inductive * 'a -> bool list array
(** Useful for [as intro_pattern] modifier *)
val compute_induction_names :
@@ -136,9 +137,6 @@ 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
-val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
-val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
-
(** Tacticals defined directly in term of Proofview *)
(** The tacticals in the module [New] are the tactical of Ltac. Their
@@ -229,7 +227,7 @@ module New : sig
val tclTIMEOUT : int -> unit tactic -> unit tactic
val tclTIME : string option -> 'a tactic -> 'a tactic
- val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> Context.Named.t
+ val nLastDecls : ('a, 'r) Proofview.Goal.t -> int -> named_context
val ifOnHyp : (identifier * types -> bool) ->
(identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) ->
@@ -238,11 +236,11 @@ module New : sig
val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic
val onLastHypId : (identifier -> unit tactic) -> unit tactic
val onLastHyp : (constr -> unit tactic) -> unit tactic
- val onLastDecl : (Context.Named.Declaration.t -> unit tactic) -> unit tactic
+ val onLastDecl : (named_declaration -> unit tactic) -> unit tactic
- val onHyps : ([ `NF ], Context.Named.t) Proofview.Goal.enter ->
- (Context.Named.t -> unit tactic) -> unit tactic
- val afterHyp : Id.t -> (Context.Named.t -> unit tactic) -> unit tactic
+ val onHyps : ([ `LZ ], named_context) Proofview.Goal.enter ->
+ (named_context -> unit tactic) -> unit tactic
+ val afterHyp : Id.t -> (named_context -> unit tactic) -> unit tactic
val tryAllHyps : (identifier -> unit tactic) -> unit tactic
val tryAllHypsAndConcl : (identifier option -> unit tactic) -> unit tactic
@@ -258,11 +256,11 @@ module New : sig
val case_then_using :
or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) ->
- constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic
+ constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic
val case_nodep_then_using :
or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) ->
- constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic
+ constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic
val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 0aab77314..e79258582 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -6,20 +6,23 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Find_subterm
open Namegen
open Declarations
open Inductiveops
open Reductionops
-open Environ
open Globnames
open Evd
open Pfedit
@@ -193,7 +196,7 @@ let introduction ?(check=true) id =
(str "Variable " ++ pr_id id ++ str " is already declared.")
in
let open Context.Named.Declaration in
- match kind_of_term (whd_evar sigma concl) with
+ match EConstr.kind sigma concl with
| Prod (_, t, b) -> unsafe_intro env store (LocalAssum (id, t)) b
| LetIn (_, c, t, b) -> unsafe_intro env store (LocalDef (id, c, t)) b
| _ -> raise (RefinerError IntroNeedsProduct)
@@ -205,7 +208,7 @@ let convert_concl ?(check=true) ty k =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- let conclty = Proofview.Goal.raw_concl gl in
+ let conclty = Proofview.Goal.concl gl in
Refine.refine ~unsafe:true { run = begin fun sigma ->
let Sigma ((), sigma, p) =
if check then begin
@@ -225,7 +228,7 @@ let convert_hyp ?(check=true) d =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let ty = Proofview.Goal.raw_concl gl in
+ let ty = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let sign = convert_hyp check (named_context_val env) sigma d in
let env = reset_with_named_context sign env in
@@ -309,24 +312,26 @@ let clear ids = clear_gen error_clear_dependency ids
let clear_for_replacing ids = clear_gen error_replacing_dependency ids
let apply_clear_request clear_flag dft c =
+ Proofview.tclEVARMAP >>= fun sigma ->
let check_isvar c =
- if not (isVar c) then
+ if not (isVar sigma c) then
error "keep/clear modifiers apply only to hypothesis names." in
let doclear = match clear_flag with
- | None -> dft && isVar c
+ | None -> dft && isVar sigma c
| Some true -> check_isvar c; true
| Some false -> false in
- if doclear then clear [destVar c]
+ if doclear then clear [destVar sigma c]
else Tacticals.New.tclIDTAC
(* Moving hypotheses *)
let move_hyp id dest =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
- let ty = Proofview.Goal.raw_concl gl in
+ let sigma = Tacmach.New.project gl in
+ let ty = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let sign = named_context_val env in
- let sign' = move_hyp_in_named_context id dest sign in
+ let sign' = move_hyp_in_named_context sigma id dest sign in
let env = reset_with_named_context sign' env in
Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar env sigma ~principal:true ~store ty
@@ -380,7 +385,7 @@ let rename_hyp repl =
in
let nhyps = List.map map hyps in
let nconcl = subst concl in
- let nctx = Environ.val_of_named_context nhyps in
+ let nctx = val_of_named_context nhyps in
let instance = List.map (NamedDecl.get_id %> mkVar) hyps in
Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance
@@ -413,7 +418,7 @@ let default_id env sigma decl =
| LocalAssum (name,t) ->
let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in
id_of_name_with_default dft name
- | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name
+ | LocalDef (name,b,_) -> id_of_name_using_hdchar env sigma b name
(* Non primitive introduction tactics are treated by intro_then_gen
There is possibly renaming, with possibly names to avoid and
@@ -497,7 +502,7 @@ fun env sigma p -> function
let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in
Sigma (arg :: rem, sigma, r)
-let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) with
+let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast sigma cl) with
| Prod (na, c1, b) ->
if Int.equal k 1 then
try
@@ -510,7 +515,7 @@ let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) w
| _ -> error "Not enough products."
(* Refine as a fixpoint *)
-let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let mutual_fix f n rest j = Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
@@ -538,7 +543,7 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl ->
let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
let typarray = Array.of_list (List.map pi3 all) in
let bodies = Array.of_list evs in
- let oterm = Term.mkFix ((indxs,0),(funnames,typarray,bodies)) in
+ let oterm = mkFix ((indxs,0),(funnames,typarray,bodies)) in
Sigma (oterm, sigma, p)
end }
end }
@@ -555,7 +560,7 @@ let fix ido n = match ido with
let rec check_is_mutcoind env sigma cl =
let b = whd_all env sigma cl in
- match kind_of_term b with
+ match EConstr.kind sigma b with
| Prod (na, c1, b) ->
let open Context.Rel.Declaration in
check_is_mutcoind (push_rel (LocalAssum (na,c1)) env) sigma b
@@ -566,7 +571,7 @@ let rec check_is_mutcoind env sigma cl =
error "All methods must construct elements in coinductive types."
(* Refine as a cofixpoint *)
-let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let mutual_cofix f others j = Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
@@ -589,7 +594,7 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl ->
let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
let typarray = Array.of_list types in
let bodies = Array.of_list evs in
- let oterm = Term.mkCoFix (0, (funnames, typarray, bodies)) in
+ let oterm = mkCoFix (0, (funnames, typarray, bodies)) in
Sigma (oterm, sigma, p)
end }
end }
@@ -612,7 +617,7 @@ type tactic_reduction = env -> evar_map -> constr -> constr
let pf_reduce_decl redfun where decl gl =
let open Context.Named.Declaration in
- let redfun' = Tacmach.New.pf_apply redfun gl in
+ let redfun' c = Tacmach.New.pf_apply redfun gl c in
match decl with
| LocalAssum (id,ty) ->
if where == InHypValueOnly then
@@ -692,12 +697,12 @@ let bind_red_expr_occurrences occs nbcl redexp =
certain hypothesis *)
let reduct_in_concl (redfun,sty) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
convert_concl_no_check (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl)) sty
end }
let reduct_in_hyp ?(check=false) redfun (id,where) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
convert_hyp ~check (pf_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl)
end }
@@ -726,14 +731,14 @@ let pf_e_reduce_decl redfun where decl gl =
Sigma (LocalDef (id, b', ty'), sigma, p +> q)
let e_reduct_in_concl ~check (redfun, sty) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in
Sigma (convert_concl ~check c' sty, sigma, p)
end }
let e_reduct_in_hyp ?(check=false) redfun (id, where) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let Sigma (decl', sigma, p) = pf_e_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl in
Sigma (convert_hyp ~check decl', sigma, p)
end }
@@ -748,7 +753,7 @@ let e_reduct_option ?(check=false) redfun = function
let e_change_in_concl (redfun,sty) =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
- let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in
+ let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.concl gl) in
Sigma (convert_concl_no_check c sty, sigma, p)
end }
@@ -777,7 +782,7 @@ let e_change_in_hyp redfun (id,where) =
Sigma (convert_hyp c, sigma, p)
end }
-type change_arg = Pattern.patvar_map -> constr Sigma.run
+type change_arg = Pattern.patvar_map -> EConstr.constr Sigma.run
let make_change_arg c pats =
{ run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma }
@@ -791,15 +796,15 @@ let check_types env sigma mayneedglobalcheck deep newc origc =
let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in
if not b then
if
- isSort (whd_all env sigma t1) &&
- isSort (whd_all env sigma t2)
+ isSort sigma (whd_all env sigma t1) &&
+ isSort sigma (whd_all env sigma t2)
then (mayneedglobalcheck := true; sigma)
else
user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.")
else sigma
end
else
- if not (isSort (whd_all env sigma t1)) then
+ if not (isSort sigma (whd_all env sigma t1)) then
user_err ~hdr:"convert-check-hyp" (str "Not a type.")
else sigma
@@ -888,7 +893,7 @@ let reduction_clause redexp cl =
let reduce redexp cl =
let trace () =
let open Printer in
- let pr = (pr_constr, pr_lconstr, pr_evaluable_reference, pr_constr_pattern) in
+ let pr = (pr_econstr, pr_leconstr, pr_evaluable_reference, pr_constr_pattern) in
Pp.(hov 2 (Pputils.pr_red_expr pr str redexp))
in
Proofview.Trace.name_tactic trace begin
@@ -936,13 +941,13 @@ let build_intro_tac id dest tac = match dest with
let rec intro_then_gen name_flag move_flag force_flag dep_flag tac =
let open Context.Rel.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
- let concl = nf_evar (Tacmach.New.project gl) concl in
- match kind_of_term concl with
- | Prod (name,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
+ match EConstr.kind sigma concl with
+ | Prod (name,t,u) when not dep_flag || not (noccurn sigma 1 u) ->
let name = find_name false (LocalAssum (name,t)) name_flag gl in
build_intro_tac name move_flag tac
- | LetIn (name,b,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
+ | LetIn (name,b,t,u) when not dep_flag || not (noccurn sigma 1 u) ->
let name = find_name false (LocalDef (name,b,t)) name_flag gl in
build_intro_tac name move_flag tac
| _ ->
@@ -1067,14 +1072,14 @@ let intros_replacing ids =
(* User-level introduction tactics *)
-let lookup_hypothesis_as_renamed env ccl = function
- | AnonHyp n -> Detyping.lookup_index_as_renamed env ccl n
- | NamedHyp id -> Detyping.lookup_name_as_displayed env ccl id
+let lookup_hypothesis_as_renamed env sigma ccl = function
+ | AnonHyp n -> Detyping.lookup_index_as_renamed env sigma ccl n
+ | NamedHyp id -> Detyping.lookup_name_as_displayed env sigma ccl id
let lookup_hypothesis_as_renamed_gen red h gl =
let env = Proofview.Goal.env gl in
let rec aux ccl =
- match lookup_hypothesis_as_renamed env ccl h with
+ match lookup_hypothesis_as_renamed env (Tacmach.New.project gl) ccl h with
| None when red ->
let (redfun, _) = Redexpr.reduction_of_red_expr env (Red true) in
let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) ccl in
@@ -1107,7 +1112,7 @@ let depth_of_quantified_hypothesis red h gl =
str".")
let intros_until_gen red h =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let n = depth_of_quantified_hypothesis red h gl in
Tacticals.New.tclDO n (if red then introf else intro)
end }
@@ -1217,13 +1222,13 @@ let cut c =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Proofview.Goal.concl gl in
let is_sort =
try
(** Backward compat: ensure that [c] is well-typed. *)
let typ = Typing.unsafe_type_of env sigma c in
let typ = whd_all env sigma typ in
- match kind_of_term typ with
+ match EConstr.kind sigma typ with
| Sort _ -> true
| _ -> false
with e when Pretype_errors.precatchable_exception e -> false
@@ -1243,6 +1248,7 @@ let cut c =
end }
let error_uninstantiated_metas t clenv =
+ let t = EConstr.Unsafe.to_constr t in
let na = meta_name clenv.evd (List.hd (Metaset.elements (metavars_of t))) in
let id = match na with Name id -> id | _ -> anomaly (Pp.str "unnamed dependent meta")
in user_err (str "Cannot find an instance for " ++ pr_id id ++ str".")
@@ -1281,7 +1287,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true)
in
let new_hyp_typ = clenv_type clenv in
if not with_evars then check_unresolved_evars_of_metas sigma0 clenv;
- if not with_evars && occur_meta new_hyp_typ then
+ if not with_evars && occur_meta clenv.evd new_hyp_typ then
error_uninstantiated_metas new_hyp_typ clenv;
let new_hyp_prf = clenv_value clenv in
let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in
@@ -1300,22 +1306,22 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true)
(* Elimination tactics *)
(********************************************)
-let last_arg c = match kind_of_term c with
+let last_arg sigma c = match EConstr.kind sigma c with
| App (f,cl) ->
Array.last cl
| _ -> anomaly (Pp.str "last_arg")
-let nth_arg i c =
- if Int.equal i (-1) then last_arg c else
- match kind_of_term c with
+let nth_arg sigma i c =
+ if Int.equal i (-1) then last_arg sigma c else
+ match EConstr.kind sigma c with
| App (f,cl) -> cl.(i)
| _ -> anomaly (Pp.str "nth_arg")
-let index_of_ind_arg t =
- let rec aux i j t = match kind_of_term t with
+let index_of_ind_arg sigma t =
+ let rec aux i j t = match EConstr.kind sigma t with
| Prod (_,t,u) ->
(* heuristic *)
- if isInd (fst (decompose_app t)) then aux (Some j) (j+1) u
+ if isInd sigma (fst (decompose_app sigma t)) then aux (Some j) (j+1) u
else aux i (j+1) u
| _ -> match i with
| Some i -> i
@@ -1330,14 +1336,14 @@ let enforce_prop_bound_names rename tac =
(* so as to avoid having hypothesis such as "t:True", "n:~A" when calling *)
(* elim or induction with schemes built by Indrec.build_induction_scheme *)
let rec aux env sigma i t =
- if i = 0 then t else match kind_of_term t with
+ if i = 0 then t else match EConstr.kind sigma t with
| Prod (Name _ as na,t,t') ->
let very_standard = true in
let na =
if Retyping.get_sort_family_of env sigma t = InProp then
(* "very_standard" says that we should have "H" names only, but
this would break compatibility even more... *)
- let s = match Namegen.head_name t with
+ let s = match Namegen.head_name sigma t with
| Some id when not very_standard -> string_of_id id
| _ -> "" in
Name (add_suffix Namegen.default_prop_ident s)
@@ -1348,9 +1354,9 @@ let enforce_prop_bound_names rename tac =
mkProd (Anonymous,t,aux (push_rel (LocalAssum (Anonymous,t)) env) sigma (i-1) t')
| LetIn (na,c,t,t') ->
mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t')
- | _ -> print_int i; Feedback.msg_notice (print_constr t); assert false in
+ | _ -> assert false in
let rename_branch i =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let t = Proofview.Goal.concl gl in
@@ -1362,10 +1368,10 @@ let enforce_prop_bound_names rename tac =
| _ ->
tac
-let rec contract_letin_in_lam_header c =
- match kind_of_term c with
- | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header c)
- | LetIn (x,b,t,c) -> contract_letin_in_lam_header (subst1 b c)
+let rec contract_letin_in_lam_header sigma c =
+ match EConstr.kind sigma c with
+ | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header sigma c)
+ | LetIn (x,b,t,c) -> contract_letin_in_lam_header sigma (subst1 b c)
| _ -> c
let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags ())
@@ -1373,10 +1379,10 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let elim = contract_letin_in_lam_header elim in
+ let elim = contract_letin_in_lam_header sigma elim in
let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in
let indmv =
- (match kind_of_term (nth_arg i elimclause.templval.rebus) with
+ (match EConstr.kind sigma (nth_arg sigma i elimclause.templval.rebus) with
| Meta mv -> mv
| _ -> user_err ~hdr:"elimination_clause"
(str "The type of elimination clause is not well-formed."))
@@ -1396,7 +1402,7 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags
type eliminator = {
elimindex : int option; (* None = find it automatically *)
elimrename : (bool * int array) option; (** None = don't rename Prop hyps with H-names *)
- elimbody : constr with_bindings
+ elimbody : EConstr.constr with_bindings
}
let general_elim_clause_gen elimtac indclause elim =
@@ -1406,7 +1412,7 @@ let general_elim_clause_gen elimtac indclause elim =
let (elimc,lbindelimc) = elim.elimbody in
let elimt = Retyping.get_type_of env sigma elimc in
let i =
- match elim.elimindex with None -> index_of_ind_arg elimt | Some i -> i in
+ match elim.elimindex with None -> index_of_ind_arg sigma elimt | Some i -> i in
elimtac elim.elimrename i (elimc, elimt, lbindelimc) indclause
end }
@@ -1428,18 +1434,20 @@ let general_elim with_evars clear_flag (c, lbindc) elim =
(* Case analysis tactics *)
let general_case_analysis_in_context with_evars clear_flag (c,lbindc) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl gl in
let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in
let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in
let sort = Tacticals.New.elimination_sort_of_goal gl in
+ let mind = on_snd (fun u -> EInstance.kind (Sigma.to_evar_map sigma) u) mind in
let Sigma (elim, sigma, p) =
- if occur_term c concl then
+ if occur_term (Sigma.to_evar_map sigma) c concl then
build_case_analysis_scheme env sigma mind true sort
else
build_case_analysis_scheme_default env sigma mind sort in
+ let elim = EConstr.of_constr elim in
let tac =
(general_elim with_evars clear_flag (c,lbindc)
{elimindex = None; elimbody = (elim,NoBindings);
@@ -1449,7 +1457,8 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) =
end }
let general_case_analysis with_evars clear_flag (c,lbindc as cx) =
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when lbindc == NoBindings ->
Tacticals.New.tclTHEN (try_intros_until_id_check id)
(general_case_analysis_in_context with_evars clear_flag cx)
@@ -1469,6 +1478,7 @@ let is_nonrec mind = (Global.lookup_mind (fst mind)).mind_finite == Decl_kinds.B
let find_ind_eliminator ind s gl =
let gr = lookup_eliminator ind s in
let evd, c = Tacmach.New.pf_apply Evd.fresh_global gl gr in
+ let c = EConstr.of_constr c in
evd, c
let find_eliminator c gl =
@@ -1502,7 +1512,8 @@ let elim_in_context with_evars clear_flag c = function
| None -> default_elim with_evars clear_flag c
let elim with_evars clear_flag (c,lbindc as cx) elim =
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when lbindc == NoBindings ->
Tacticals.New.tclTHEN (try_intros_until_id_check id)
(elim_in_context with_evars clear_flag cx elim)
@@ -1536,9 +1547,9 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ())
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let elim = contract_letin_in_lam_header elim in
+ let elim = contract_letin_in_lam_header sigma elim in
let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in
- let indmv = destMeta (nth_arg i elimclause.templval.rebus) in
+ let indmv = destMeta sigma (nth_arg sigma i elimclause.templval.rebus) in
let hypmv =
try match List.remove Int.equal indmv (clenv_independent elimclause) with
| [a] -> a
@@ -1551,7 +1562,7 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ())
let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in
let elimclause'' = clenv_fchain_in id ~flags hypmv elimclause' hypclause in
let new_hyp_typ = clenv_type elimclause'' in
- if Term.eq_constr hyp_typ new_hyp_typ then
+ if EConstr.eq_constr sigma hyp_typ new_hyp_typ then
user_err ~hdr:"general_rewrite_in"
(str "Nothing to rewrite in " ++ pr_id id ++ str".");
clenv_refine_in with_evars id id sigma elimclause''
@@ -1576,21 +1587,23 @@ let make_projection env sigma params cstr sign elim i n c u =
let elim = match elim with
| NotADefinedRecordUseScheme elim ->
(* bugs: goes from right to left when i increases! *)
- let decl = List.nth cstr.cs_args i in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.cs_args in
+ let decl = List.nth cs_args i in
let t = RelDecl.get_type decl in
let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> b in
- let branch = it_mkLambda_or_LetIn b cstr.cs_args in
+ let branch = it_mkLambda_or_LetIn b cs_args in
if
(* excludes dependent projection types *)
- noccur_between 1 (n-i-1) t
+ noccur_between sigma 1 (n-i-1) t
(* to avoid surprising unifications, excludes flexible
projection types or lambda which will be instantiated by Meta/Evar *)
- && not (isEvar (fst (whd_betaiota_stack sigma t)))
- && (accept_universal_lemma_under_conjunctions () || not (isRel t))
+ && not (isEvar sigma (fst (whd_betaiota_stack sigma t)))
+ && (accept_universal_lemma_under_conjunctions () || not (isRel sigma t))
then
let t = lift (i+1-n) t in
- let abselim = beta_applist (elim,params@[t;branch]) in
- let c = beta_applist (abselim, [mkApp (c, Context.Rel.to_extended_vect 0 sign)]) in
+ let abselim = beta_applist sigma (elim, params@[t;branch]) in
+ let args = Context.Rel.to_extended_vect mkRel 0 sign in
+ let c = beta_applist sigma (abselim, [mkApp (c, args)]) in
Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign)
else
None
@@ -1598,7 +1611,7 @@ let make_projection env sigma params cstr sign elim i n c u =
(* goes from left to right when i increases! *)
match List.nth l i with
| Some proj ->
- let args = Context.Rel.to_extended_vect 0 sign in
+ let args = Context.Rel.to_extended_vect mkRel 0 sign in
let proj =
if Environ.is_projection proj env then
mkProj (Projection.make proj false, mkApp (c, args))
@@ -1613,25 +1626,28 @@ let make_projection env sigma params cstr sign elim i n c u =
in elim
let descend_in_conjunctions avoid tac (err, info) c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
try
let t = Retyping.get_type_of env sigma c in
let ((ind,u),t) = reduce_to_quantified_ind env sigma t in
- let sign,ccl = decompose_prod_assum t in
- match match_with_tuple ccl with
+ let sign,ccl = EConstr.decompose_prod_assum sigma t in
+ match match_with_tuple sigma ccl with
| Some (_,_,isrec) ->
let n = (constructors_nrealargs ind).(0) in
let sort = Tacticals.New.elimination_sort_of_goal gl in
let IndType (indf,_) = find_rectype env sigma ccl in
let (_,inst), params = dest_ind_family indf in
+ let params = List.map EConstr.of_constr params in
let cstr = (get_constructors env indf).(0) in
let elim =
try DefinedRecord (Recordops.lookup_projections ind)
with Not_found ->
+ let u = EInstance.kind sigma u in
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (elim, _, _) = build_case_analysis_scheme env sigma (ind,u) false sort in
+ let elim = EConstr.of_constr elim in
NotADefinedRecordUseScheme elim in
Tacticals.New.tclORELSE0
(Tacticals.New.tclFIRST
@@ -1658,7 +1674,7 @@ let descend_in_conjunctions avoid tac (err, info) c =
(****************************************************)
let solve_remaining_apply_goals =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
if !apply_solve_class_goals then
try
@@ -1682,15 +1698,16 @@ let tclORELSEOPT t k =
Proofview.tclZERO ~info e
| Some tac -> tac)
-let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : EConstr.constr with_bindings)) =
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
let flags =
if with_delta then default_unify_flags () else default_no_delta_unify_flags () in
(* The actual type of the theorem. It will be matched against the
goal. If this fails, then the head constant will be unfolded step by
step. *)
- let concl_nprod = nb_prod_modulo_zeta concl in
+ let concl_nprod = nb_prod_modulo_zeta sigma concl in
let rec try_main_apply with_destruct c =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
@@ -1699,7 +1716,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind))
let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in
let try_apply thm_ty nprod =
try
- let n = nb_prod_modulo_zeta thm_ty - nprod in
+ let n = nb_prod_modulo_zeta sigma thm_ty - nprod in
if n<0 then error "Applied theorem has not enough premisses.";
let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in
Clenvtac.res_pf clause ~with_evars ~flags
@@ -1826,9 +1843,9 @@ let progress_with_clause flags innerclause clause =
let explain_unable_to_apply_lemma loc env sigma thm innerclause =
user_err ~loc (hov 0
(Pp.str "Unable to apply lemma of type" ++ brk(1,1) ++
- Pp.quote (Printer.pr_lconstr_env env sigma thm) ++ spc() ++
+ Pp.quote (Printer.pr_leconstr_env env sigma thm) ++ spc() ++
str "on hypothesis of type" ++ brk(1,1) ++
- Pp.quote (Printer.pr_lconstr_env innerclause.env innerclause.evd (clenv_type innerclause)) ++
+ Pp.quote (Printer.pr_leconstr_env innerclause.env innerclause.evd (clenv_type innerclause)) ++
str "."))
let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
@@ -1848,7 +1865,7 @@ let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
let apply_in_once sidecond_first with_delta with_destruct with_evars naming
id (clear_flag,(loc,(d,lbind))) tac =
let open Context.Rel.Declaration in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let flags =
@@ -1909,9 +1926,10 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam
*)
let cut_and_apply c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- match kind_of_term (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with
- | Prod (_,c1,c2) when not (dependent (mkRel 1) c2) ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
+ match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with
+ | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 ->
let concl = Proofview.Goal.concl gl in
let env = Tacmach.New.pf_env gl in
Refine.refine { run = begin fun sigma ->
@@ -1954,7 +1972,7 @@ let exact_check c =
let cast_no_check cast c =
Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
- exact_no_check (Term.mkCast (c, cast, concl))
+ exact_no_check (mkCast (c, cast, concl))
end }
let vm_cast_no_check c = cast_no_check Term.VMcast c
@@ -1962,10 +1980,11 @@ let native_cast_no_check c = cast_no_check Term.NATIVEcast c
let exact_proof c =
let open Tacmach.New in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Refine.refine { run = begin fun sigma ->
let sigma = Sigma.to_evar_map sigma in
let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in
+ let c = EConstr.of_constr c in
let sigma = Evd.merge_universe_context sigma ctx in
Sigma.Unsafe.of_pair (c, sigma)
end }
@@ -1983,7 +2002,7 @@ let assumption =
let concl = Proofview.Goal.concl gl in
let sigma = Tacmach.New.project gl in
let (sigma, is_same_type) =
- if only_eq then (sigma, Constr.equal t concl)
+ if only_eq then (sigma, EConstr.eq_constr sigma t concl)
else
let env = Proofview.Goal.env gl in
infer_conv env sigma t concl
@@ -1997,7 +2016,7 @@ let assumption =
let hyps = Proofview.Goal.hyps gl in
arec gl true hyps
end } in
- Proofview.Goal.nf_enter assumption_tac
+ Proofview.Goal.enter assumption_tac
(*****************************************************************)
(* Modification of a local context *)
@@ -2058,7 +2077,7 @@ let clear_body ids =
(** Do no recheck hypotheses that do not depend *)
let sigma =
if not seen then sigma
- else if List.exists (fun id -> occur_var_in_decl env id decl) ids then
+ else if List.exists (fun id -> occur_var_in_decl env sigma id decl) ids then
check_decl env sigma decl
else sigma
in
@@ -2067,7 +2086,7 @@ let clear_body ids =
in
let (env, sigma, _) = List.fold_left check (base_env, sigma, false) (List.rev ctx) in
let sigma =
- if List.exists (fun id -> occur_var env id concl) ids then
+ if List.exists (fun id -> occur_var env sigma id concl) ids then
check_is_type env sigma concl
else sigma
in
@@ -2102,15 +2121,17 @@ let rec intros_clearing = function
(* Keeping only a few hypotheses *)
let keep hyps =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Proofview.tclENV >>= fun env ->
let ccl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
let cl,_ =
fold_named_context_reverse (fun (clear,keep) decl ->
+ let decl = map_named_decl EConstr.of_constr decl in
let hyp = NamedDecl.get_id decl in
if Id.List.mem hyp hyps
- || List.exists (occur_var_in_decl env hyp) keep
- || occur_var env hyp ccl
+ || List.exists (occur_var_in_decl env sigma hyp) keep
+ || occur_var env sigma hyp ccl
then (clear,decl::keep)
else (hyp::clear,keep))
~init:([],[]) (Proofview.Goal.env gl)
@@ -2148,9 +2169,9 @@ let bring_hyps hyps =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in
- let args = Array.of_list (Context.Named.to_instance hyps) in
+ let args = Array.of_list (Context.Named.to_instance mkVar hyps) in
Refine.refine { run = begin fun sigma ->
let Sigma (ev, sigma, p) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
@@ -2182,7 +2203,7 @@ let check_number_of_constructors expctdnumopt i nconstr =
let constructor_tac with_evars expctdnumopt i lbind =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
- let cl = Tacmach.New.pf_nf_concl gl in
+ let cl = Tacmach.New.pf_concl gl in
let reduce_to_quantified_ind =
Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl
in
@@ -2191,9 +2212,9 @@ let constructor_tac with_evars expctdnumopt i lbind =
Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in
check_number_of_constructors expctdnumopt i nconstr;
- let Sigma (cons, sigma, p) = Sigma.fresh_constructor_instance
+ let Sigma ((cons, u), sigma, p) = Sigma.fresh_constructor_instance
(Proofview.Goal.env gl) sigma (fst mind, i) in
- let cons = mkConstructU cons in
+ let cons = mkConstructU (cons, EInstance.make u) in
let apply_tac = general_apply true false with_evars None (dloc,(cons,lbind)) in
let tac =
@@ -2221,7 +2242,7 @@ let any_constructor with_evars tacopt =
let t = match tacopt with None -> Proofview.tclUNIT () | Some t -> t in
let tac i = Tacticals.New.tclTHEN (constructor_tac with_evars None i NoBindings) t in
Proofview.Goal.enter { enter = begin fun gl ->
- let cl = Tacmach.New.pf_nf_concl gl in
+ let cl = Tacmach.New.pf_concl gl in
let reduce_to_quantified_ind =
Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl
in
@@ -2281,7 +2302,7 @@ let my_find_eq_data_decompose gl t =
| Constr_matching.PatternMatchingFailure -> None
let intro_decomp_eq loc l thin tac id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let c = mkVar id in
let t = Tacmach.New.pf_unsafe_type_of gl c in
let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in
@@ -2319,24 +2340,25 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac =
List.filter (fun (_,id) -> not (Id.equal id id')) thin in
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let type_of = Tacmach.New.pf_unsafe_type_of gl in
let whd_all = Tacmach.New.pf_apply whd_all gl in
let t = whd_all (type_of (mkVar id)) in
- let eqtac, thin = match match_with_equality_type t with
+ let eqtac, thin = match match_with_equality_type sigma t with
| Some (hdcncl,[_;lhs;rhs]) ->
- if l2r && isVar lhs && not (occur_var env (destVar lhs) rhs) then
- let id' = destVar lhs in
+ if l2r && isVar sigma lhs && not (occur_var env sigma (destVar sigma lhs) rhs) then
+ let id' = destVar sigma lhs in
subst_on l2r id' rhs, early_clear id' thin
- else if not l2r && isVar rhs && not (occur_var env (destVar rhs) lhs) then
- let id' = destVar rhs in
+ else if not l2r && isVar sigma rhs && not (occur_var env sigma (destVar sigma rhs) lhs) then
+ let id' = destVar sigma rhs in
subst_on l2r id' lhs, early_clear id' thin
else
Tacticals.New.tclTHEN (rew_on l2r onConcl) (clear [id]),
thin
| Some (hdcncl,[c]) ->
let l2r = not l2r in (* equality of the form eq_true *)
- if isVar c then
- let id' = destVar c in
+ if isVar sigma c then
+ let id' = destVar sigma c in
Tacticals.New.tclTHEN (rew_on l2r allHypsAndConcl)
(clear_var_and_eq id'),
early_clear id' thin
@@ -2535,9 +2557,9 @@ let ipat_of_name = function
| Anonymous -> None
| Name id -> Some (dloc, IntroNaming (IntroIdentifier id))
-let head_ident c =
- let c = fst (decompose_app ((strip_lam_assum c))) in
- if isVar c then Some (destVar c) else None
+let head_ident sigma c =
+ let c = fst (decompose_app sigma (snd (decompose_lam_assum sigma c))) in
+ if isVar sigma c then Some (destVar sigma c) else None
let assert_as first hd ipat t =
let naming,tac = prepare_intros false IntroAnonymous MoveLast ipat in
@@ -2619,7 +2641,9 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty =
let eqdata = build_coq_eq_data () in
let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in
let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in
+ let eq = EConstr.of_constr eq in
let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in
+ let refl = EConstr.of_constr refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in
@@ -2650,6 +2674,7 @@ let insert_before decls lasthyp env =
| Some id ->
Environ.fold_named_context
(fun _ d env ->
+ let d = map_named_decl EConstr.of_constr d in
let env = if Id.equal id (NamedDecl.get_id d) then push_named_context decls env else env in
push_named d env)
~init:(reset_context env) env
@@ -2674,7 +2699,9 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
let eqdata = build_coq_eq_data () in
let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in
let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in
+ let eq = EConstr.of_constr eq in
let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in
+ let refl = EConstr.of_constr refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in
@@ -2686,7 +2713,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
Sigma (mkNamedLetIn id c t x, sigma, p)
let letin_tac with_eq id c ty occs =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let ccl = Proofview.Goal.concl gl in
@@ -2703,7 +2730,7 @@ let letin_tac with_eq id c ty occs =
end }
let letin_pat_tac with_eq id c occs =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let ccl = Proofview.Goal.concl gl in
@@ -2725,7 +2752,8 @@ let forward b usetac ipat c =
| None ->
Proofview.Goal.enter { enter = begin fun gl ->
let t = Tacmach.New.pf_get_type_of gl c in
- let hd = head_ident c in
+ let sigma = Tacmach.New.project gl in
+ let hd = head_ident sigma c in
Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c)
end }
| Some tac ->
@@ -2748,22 +2776,22 @@ let enough_by na t tac = forward false (Some (Some tac)) (ipat_of_name na) t
(* Compute a name for a generalization *)
-let generalized_name c t ids cl = function
+let generalized_name env sigma c t ids cl = function
| Name id as na ->
if Id.List.mem id ids then
user_err (pr_id id ++ str " is already used.");
na
| Anonymous ->
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id ->
(* Keep the name even if not occurring: may be used by intros later *)
Name id
| _ ->
- if noccurn 1 cl then Anonymous else
+ if noccurn sigma 1 cl then Anonymous else
(* On ne s'etait pas casse la tete : on avait pris pour nom de
variable la premiere lettre du type, meme si "c" avait ete une
constante dont on aurait pu prendre directement le nom *)
- named_hd (Global.env()) t Anonymous
+ named_hd env sigma t Anonymous
(* Abstract over [c] in [forall x1:A1(c)..xi:Ai(c).T(c)] producing
[forall x, x1:A1(x1), .., xi:Ai(x). T(x)] with all [c] abtracted in [Ai]
@@ -2771,11 +2799,11 @@ let generalized_name c t ids cl = function
let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl =
let open Context.Rel.Declaration in
- let decls,cl = decompose_prod_n_assum i cl in
+ let decls,cl = decompose_prod_n_assum sigma i cl in
let dummy_prod = it_mkProd_or_LetIn mkProp decls in
- let newdecls,_ = decompose_prod_n_assum i (subst_term_gen eq_constr_nounivs c dummy_prod) in
+ let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in
let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in
- let na = generalized_name c t ids cl' na in
+ let na = generalized_name env sigma c t ids cl' na in
let decl = match b with
| None -> LocalAssum (na,t)
| Some b -> LocalDef (na,b,t)
@@ -2788,13 +2816,20 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
let sigma, t = Typing.type_of env sigma c in
generalize_goal_gen env sigma ids i o t cl
+let new_generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
+ let env = Tacmach.New.pf_env gl in
+ let ids = Tacmach.New.pf_ids_of_hyps gl in
+ let sigma, t = Typing.type_of env sigma c in
+ generalize_goal_gen env sigma ids i o t cl
+
let old_generalize_dep ?(with_let=false) c gl =
let env = pf_env gl in
let sign = pf_hyps gl in
+ let sigma = project gl in
let init_ids = ids_of_named_context (Global.named_context()) in
- let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) =
- if List.exists (fun d' -> occur_var_in_decl env (NamedDecl.get_id d') d) toquant
- || dependent_in_decl c d then
+ let seek (d:named_declaration) (toquant:named_context) =
+ if List.exists (fun d' -> occur_var_in_decl env sigma (NamedDecl.get_id d') d) toquant
+ || dependent_in_decl sigma c d then
d::toquant
else
toquant in
@@ -2803,7 +2838,7 @@ let old_generalize_dep ?(with_let=false) c gl =
let qhyps = List.map NamedDecl.get_id to_quantify_rev in
let tothin = List.filter (fun id -> not (Id.List.mem id init_ids)) qhyps in
let tothin' =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id when mem_named_context_val id (val_of_named_context sign) && not (Id.List.mem id init_ids)
-> id::tothin
| _ -> tothin
@@ -2811,7 +2846,7 @@ let old_generalize_dep ?(with_let=false) c gl =
let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in
let body =
if with_let then
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id -> id |> Tacmach.pf_get_hyp gl |> NamedDecl.get_value
| _ -> None
else None
@@ -2820,7 +2855,7 @@ let old_generalize_dep ?(with_let=false) c gl =
(cl',project gl) in
(** Check that the generalization is indeed well-typed *)
let (evd, _) = Typing.type_of env evd cl'' in
- let args = Context.Named.to_instance to_quantify_rev in
+ let args = Context.Named.to_instance mkVar to_quantify_rev in
tclTHENLIST
[tclEVARS evd;
Proofview.V82.of_tactic (apply_type cl'' (if Option.is_empty body then c::args else args));
@@ -2831,10 +2866,10 @@ let generalize_dep ?(with_let = false) c =
Proofview.V82.tactic (old_generalize_dep ~with_let c)
(** *)
-let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+let generalize_gen_let lconstr = Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let newcl, evd =
- List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr
+ List.fold_right_i (new_generalize_goal gl) 0 lconstr
(Tacmach.New.pf_concl gl,Tacmach.New.project gl)
in
let (evd, _) = Typing.type_of env evd newcl in
@@ -2910,19 +2945,19 @@ let specialize (c,lbind) ipat =
let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in
let rec chk = function
| [] -> []
- | t::l -> if occur_meta t then [] else t :: chk l
+ | t::l -> if occur_meta clause.evd t then [] else t :: chk l
in
let tstack = chk tstack in
let term = applist(thd,List.map (nf_evar clause.evd) tstack) in
- if occur_meta term then
+ if occur_meta clause.evd term then
user_err (str "Cannot infer an instance for " ++
- pr_name (meta_name clause.evd (List.hd (collect_metas term))) ++
+ pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd term))) ++
str ".");
clause.evd, term in
let typ = Retyping.get_type_of env sigma term in
let tac =
- match kind_of_term (fst(decompose_app (snd(decompose_lam_assum c)))) with
+ match EConstr.kind sigma (fst(EConstr.decompose_app sigma (snd(EConstr.decompose_lam_assum sigma c)))) with
| Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) ->
(* Like assert (id:=id args) but with the concept of specialization *)
let naming,tac =
@@ -2964,6 +2999,7 @@ let unfold_body x =
(pr_id x ++ str" is not a defined hypothesis.")
| LocalDef (_,xval,_) -> xval
in
+ let xval = EConstr.of_constr xval in
Tacticals.New.afterHyp x begin fun aft ->
let hl = List.fold_right (fun decl cl -> (NamedDecl.get_id decl, InHyp) :: cl) aft [] in
let rfun _ _ c = replace_vars [x, xval] c in
@@ -3016,7 +3052,7 @@ let warn_unused_intro_pattern =
strbrk"Unused introduction " ++ str (String.plural (List.length names) "pattern")
++ str": " ++ prlist_with_sep spc
(Miscprint.pr_intro_pattern
- (fun c -> Printer.pr_constr (fst (run_delayed (Global.env()) Evd.empty c)))) names)
+ (fun c -> Printer.pr_econstr (fst (run_delayed (Global.env()) Evd.empty c)))) names)
let check_unused_names names =
if not (List.is_empty names) then
@@ -3150,12 +3186,12 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names =
substitutions aussi sur l'argument voisin *)
let expand_projections env sigma c =
- let sigma = Sigma.to_evar_map sigma in
let rec aux env c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) []
- | _ -> map_constr_with_full_binders push_rel aux env c
- in aux env c
+ | _ -> map_constr_with_full_binders sigma push_rel aux env c
+ in
+ aux env c
(* Marche pas... faut prendre en compte l'occurrence précise... *)
@@ -3163,13 +3199,13 @@ let expand_projections env sigma c =
let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 (Proofview.Goal.assume gl) in
let reduce_to_quantified_ref = Tacmach.New.pf_apply reduce_to_quantified_ref gl in
let typ0 = reduce_to_quantified_ref indref tmptyp0 in
- let prods, indtyp = decompose_prod_assum typ0 in
- let hd,argl = decompose_app indtyp in
+ let prods, indtyp = decompose_prod_assum sigma typ0 in
+ let hd,argl = decompose_app sigma indtyp in
let env' = push_rel_context prods env in
- let sigma = Proofview.Goal.sigma gl in
let params = List.firstn nparams argl in
let params' = List.map (expand_projections env' sigma) params in
(* le gl est important pour ne pas préévaluer *)
@@ -3181,17 +3217,18 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
(tac avoid)
else
let c = List.nth argl (i-1) in
- match kind_of_term c with
- | Var id when not (List.exists (occur_var env id) args') &&
- not (List.exists (occur_var env id) params') ->
+ match EConstr.kind sigma c with
+ | Var id when not (List.exists (fun c -> occur_var env sigma id c) args') &&
+ not (List.exists (fun c -> occur_var env sigma id c) params') ->
(* Based on the knowledge given by the user, all
constraints on the variable are generalizable in the
current environment so that it is clearable after destruction *)
atomize_one (i-1) (c::args) (c::args') (id::avoid)
| _ ->
let c' = expand_projections env' sigma c in
- if List.exists (dependent c) params' ||
- List.exists (dependent c) args'
+ let dependent t = dependent sigma c t in
+ if List.exists dependent params' ||
+ List.exists dependent args'
then
(* This is a case where the argument is constrained in a
way which would require some kind of inversion; we
@@ -3203,11 +3240,11 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
(* We reason blindly on the term and do as if it were
generalizable, ignoring the constraints coming from
its structure *)
- let id = match kind_of_term c with
+ let id = match EConstr.kind sigma c with
| Var id -> id
| _ ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
- id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in
+ id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in
let x = fresh_id_in_env avoid id env in
Tacticals.New.tclTHEN
(letin_tac None (Name x) c None allHypsAndConcl)
@@ -3281,7 +3318,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
exception Shunt of Id.t move_location
-let cook_sign hyp0_opt inhyps indvars env =
+let cook_sign hyp0_opt inhyps indvars env sigma =
(* First phase from L to R: get [toclear], [decldep] and [statuslist]
for the hypotheses before (= more ancient than) hyp0 (see above) *)
let toclear = ref [] in
@@ -3293,6 +3330,7 @@ let cook_sign hyp0_opt inhyps indvars env =
let before = ref true in
let maindep = ref false in
let seek_deps env decl rhyp =
+ let decl = map_named_decl EConstr.of_constr decl in
let hyp = NamedDecl.get_id decl in
if (match hyp0_opt with Some hyp0 -> Id.equal hyp hyp0 | _ -> false)
then begin
@@ -3308,11 +3346,11 @@ let cook_sign hyp0_opt inhyps indvars env =
rhyp
end else
let dephyp0 = List.is_empty inhyps &&
- (Option.cata (fun id -> occur_var_in_decl env id decl) false hyp0_opt)
+ (Option.cata (fun id -> occur_var_in_decl env sigma id decl) false hyp0_opt)
in
let depother = List.is_empty inhyps &&
- (List.exists (fun id -> occur_var_in_decl env id decl) indvars ||
- List.exists (fun decl' -> occur_var_in_decl env (NamedDecl.get_id decl') decl) !decldeps)
+ (List.exists (fun id -> occur_var_in_decl env sigma id decl) indvars ||
+ List.exists (fun decl' -> occur_var_in_decl env sigma (NamedDecl.get_id decl') decl) !decldeps)
in
if not (List.is_empty inhyps) && Id.List.mem hyp inhyps
|| dephyp0 || depother
@@ -3385,15 +3423,15 @@ type elim_scheme = {
elimc: constr with_bindings option;
elimt: types;
indref: global_reference option;
- params: Context.Rel.t; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
+ params: rel_context; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
nparams: int; (* number of parameters *)
- predicates: Context.Rel.t; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
+ predicates: rel_context; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
npredicates: int; (* Number of predicates *)
- branches: Context.Rel.t; (* branchr,...,branch1 *)
+ branches: rel_context; (* branchr,...,branch1 *)
nbranches: int; (* Number of branches *)
- args: Context.Rel.t; (* (xni, Ti_ni) ... (x1, Ti_1) *)
+ args: rel_context; (* (xni, Ti_ni) ... (x1, Ti_1) *)
nargs: int; (* number of arguments *)
- indarg: Context.Rel.Declaration.t option; (* Some (H,I prm1..prmp x1...xni)
+ indarg: rel_declaration option; (* Some (H,I prm1..prmp x1...xni)
if HI is in premisses, None otherwise *)
concl: types; (* Qi x1...xni HI (f...), HI and (f...)
are optional and mutually exclusive *)
@@ -3458,13 +3496,13 @@ let error_ind_scheme s =
let s = if not (String.is_empty s) then s^" " else s in
user_err ~hdr:"Tactics" (str "Cannot recognize " ++ str s ++ str "an induction scheme.")
-let glob = Universes.constr_of_global
+let glob c = EConstr.of_constr (Universes.constr_of_global c)
let coq_eq = lazy (glob (Coqlib.build_coq_eq ()))
let coq_eq_refl = lazy (glob (Coqlib.build_coq_eq_refl ()))
-let coq_heq = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq")
-let coq_heq_refl = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl")
+let coq_heq = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq"))
+let coq_heq_refl = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl"))
let mkEq t x y =
@@ -3491,24 +3529,24 @@ let lift_togethern n l =
let lift_list l = List.map (lift 1) l
-let ids_of_constr ?(all=false) vars c =
+let ids_of_constr sigma ?(all=false) vars c =
let rec aux vars c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id -> Id.Set.add id vars
| App (f, args) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Construct ((ind,_),_)
| Ind (ind,_) ->
let (mib,mip) = Global.lookup_inductive ind in
Array.fold_left_from
(if all then 0 else mib.Declarations.mind_nparams)
aux vars args
- | _ -> Term.fold_constr aux vars c)
- | _ -> Term.fold_constr aux vars c
+ | _ -> EConstr.fold sigma aux vars c)
+ | _ -> EConstr.fold sigma aux vars c
in aux vars c
-let decompose_indapp f args =
- match kind_of_term f with
+let decompose_indapp sigma f args =
+ match EConstr.kind sigma f with
| Construct ((ind,_),_)
| Ind (ind,_) ->
let (mib,mip) = Global.lookup_inductive ind in
@@ -3558,7 +3596,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
Sigma (mkApp (appeqs, abshypt), sigma, p)
end }
-let hyps_of_vars env sign nogen hyps =
+let hyps_of_vars env sigma sign nogen hyps =
if Id.Set.is_empty hyps then []
else
let (_,lh) =
@@ -3568,7 +3606,7 @@ let hyps_of_vars env sign nogen hyps =
if Id.Set.mem x nogen then (hs,hl)
else if Id.Set.mem x hs then (hs,x::hl)
else
- let xvars = global_vars_set_of_decl env d in
+ let xvars = global_vars_set_of_decl env sigma d in
if not (Id.Set.is_empty (Id.Set.diff xvars hs)) then
(Id.Set.add x hs, x :: hl)
else (hs, hl))
@@ -3578,11 +3616,11 @@ let hyps_of_vars env sign nogen hyps =
exception Seen
-let linear vars args =
+let linear sigma vars args =
let seen = ref vars in
try
Array.iter (fun i ->
- let rels = ids_of_constr ~all:true Id.Set.empty i in
+ let rels = ids_of_constr sigma ~all:true Id.Set.empty i in
let seen' =
Id.Set.fold (fun id acc ->
if Id.Set.mem id acc then raise Seen
@@ -3597,14 +3635,15 @@ let is_defined_variable env id =
env |> lookup_named id |> is_local_def
let abstract_args gl generalize_vars dep id defined f args =
+ let open Tacmach.New in
let open Context.Rel.Declaration in
- let sigma = ref (Tacmach.project gl) in
- let env = Tacmach.pf_env gl in
- let concl = Tacmach.pf_concl gl in
- let dep = dep || dependent (mkVar id) concl in
+ let sigma = ref (Tacmach.New.project gl) in
+ let env = Tacmach.New.pf_env gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let dep = dep || local_occur_var !sigma id concl in
let avoid = ref [] in
let get_id name =
- let id = fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in
+ let id = new_fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in
avoid := id :: !avoid; id
in
(* Build application generalized w.r.t. the argument plus the necessary eqs.
@@ -3619,13 +3658,13 @@ let abstract_args gl generalize_vars dep id defined f args =
let decl = List.hd rel in
RelDecl.get_name decl, RelDecl.get_type decl, c
in
- let argty = Tacmach.pf_unsafe_type_of gl arg in
+ let argty = Tacmach.New.pf_unsafe_type_of gl arg in
let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in
let () = sigma := sigma' in
let lenctx = List.length ctx in
let liftargty = lift lenctx argty in
- let leq = constr_cmp Reduction.CUMUL liftargty ty in
- match kind_of_term arg with
+ let leq = constr_cmp !sigma Reduction.CUMUL liftargty ty in
+ match EConstr.kind !sigma arg with
| Var id when not (is_defined_variable env id) && leq && not (Id.Set.mem id nongenvars) ->
(subst1 arg arity, ctx, ctxenv, mkApp (c, [|arg|]), args, eqs, refls,
Id.Set.add id nongenvars, Id.Set.remove id vars, env)
@@ -3644,23 +3683,23 @@ let abstract_args gl generalize_vars dep id defined f args =
in
let eqs = eq :: lift_list eqs in
let refls = refl :: refls in
- let argvars = ids_of_constr vars arg in
+ let argvars = ids_of_constr !sigma vars arg in
(arity, ctx, push_rel decl ctxenv, c', args, eqs, refls,
nongenvars, Id.Set.union argvars vars, env)
in
- let f', args' = decompose_indapp f args in
+ let f', args' = decompose_indapp !sigma f args in
let dogen, f', args' =
- let parvars = ids_of_constr ~all:true Id.Set.empty f' in
- if not (linear parvars args') then true, f, args
+ let parvars = ids_of_constr !sigma ~all:true Id.Set.empty f' in
+ if not (linear !sigma parvars args') then true, f, args
else
- match Array.findi (fun i x -> not (isVar x) || is_defined_variable env (destVar x)) args' with
+ match Array.findi (fun i x -> not (isVar !sigma x) || is_defined_variable env (destVar !sigma x)) args' with
| None -> false, f', args'
| Some nonvar ->
let before, after = Array.chop nonvar args' in
true, mkApp (f', before), after
in
if dogen then
- let tyf' = Tacmach.pf_unsafe_type_of gl f' in
+ let tyf' = Tacmach.New.pf_unsafe_type_of gl f' in
let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env =
Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args'
in
@@ -3668,14 +3707,14 @@ let abstract_args gl generalize_vars dep id defined f args =
let vars =
if generalize_vars then
let nogen = Id.Set.add id nogen in
- hyps_of_vars (pf_env gl) (pf_hyps gl) nogen vars
+ hyps_of_vars (pf_env gl) (project gl) (Proofview.Goal.hyps gl) nogen vars
else []
in
let body, c' =
if defined then Some c', Retyping.get_type_of ctxenv !sigma c'
else None, c'
in
- let typ = Tacmach.pf_get_hyp_typ gl id in
+ let typ = Tacmach.New.pf_get_hyp_typ id gl in
let tac = make_abstract_generalize (pf_env gl) id typ concl dep ctx body c' eqs args refls in
let tac = Proofview.Unsafe.tclEVARS !sigma <*> tac in
Some (tac, dep, succ (List.length ctx), vars)
@@ -3683,21 +3722,22 @@ let abstract_args gl generalize_vars dep id defined f args =
let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
let open Context.Named.Declaration in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Coqlib.check_required_library Coqlib.jmeq_module_name;
+ let sigma = Tacmach.New.project gl in
let (f, args, def, id, oldid) =
let oldid = Tacmach.New.pf_get_new_id id gl in
match Tacmach.New.pf_get_hyp id gl with
- | LocalAssum (_,t) -> let f, args = decompose_app t in
+ | LocalAssum (_,t) -> let f, args = decompose_app sigma t in
(f, args, false, id, oldid)
| LocalDef (_,t,_) ->
- let f, args = decompose_app t in
+ let f, args = decompose_app sigma t in
(f, args, true, id, oldid)
in
if List.is_empty args then Proofview.tclUNIT ()
else
let args = Array.of_list args in
- let newc = Tacmach.New.of_old (fun gl -> abstract_args gl generalize_vars force_dep id def f args) gl in
+ let newc = abstract_args gl generalize_vars force_dep id def f args in
match newc with
| None -> Proofview.tclUNIT ()
| Some (tac, dep, n, vars) ->
@@ -3720,9 +3760,12 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
Tacticals.New.tclTRY (generalize_dep ~with_let:true (mkVar id))) vars])
end }
-let rec compare_upto_variables x y =
- if (isVar x || isRel x) && (isVar y || isRel y) then true
- else compare_constr compare_upto_variables x y
+let compare_upto_variables sigma x y =
+ let rec compare x y =
+ if (isVar sigma x || isRel sigma x) && (isVar sigma y || isRel sigma y) then true
+ else compare_constr sigma compare x y
+ in
+ compare x y
let specialize_eqs id gl =
let open Context.Rel.Declaration in
@@ -3730,21 +3773,21 @@ let specialize_eqs id gl =
let ty = Tacmach.pf_get_hyp_typ gl id in
let evars = ref (project gl) in
let unif env evars c1 c2 =
- compare_upto_variables c1 c2 && Evarconv.e_conv env evars c1 c2
+ compare_upto_variables !evars c1 c2 && Evarconv.e_conv env evars c1 c2
in
let rec aux in_eqs ctx acc ty =
- match kind_of_term ty with
+ match EConstr.kind !evars ty with
| Prod (na, t, b) ->
- (match kind_of_term t with
- | App (eq, [| eqty; x; y |]) when Term.eq_constr (Lazy.force coq_eq) eq ->
- let c = if noccur_between 1 (List.length ctx) x then y else x in
+ (match EConstr.kind !evars t with
+ | App (eq, [| eqty; x; y |]) when EConstr.eq_constr !evars (Lazy.force coq_eq) eq ->
+ let c = if noccur_between !evars 1 (List.length ctx) x then y else x in
let pt = mkApp (Lazy.force coq_eq, [| eqty; c; c |]) in
let p = mkApp (Lazy.force coq_eq_refl, [| eqty; c |]) in
if unif (push_rel_context ctx env) evars pt t then
aux true ctx (mkApp (acc, [| p |])) (subst1 p b)
else acc, in_eqs, ctx, ty
- | App (heq, [| eqty; x; eqty'; y |]) when Term.eq_constr heq (Lazy.force coq_heq) ->
- let eqt, c = if noccur_between 1 (List.length ctx) x then eqty', y else eqty, x in
+ | App (heq, [| eqty; x; eqty'; y |]) when EConstr.eq_constr !evars heq (Lazy.force coq_heq) ->
+ let eqt, c = if noccur_between !evars 1 (List.length ctx) x then eqty', y else eqty, x in
let pt = mkApp (Lazy.force coq_heq, [| eqt; c; eqt; c |]) in
let p = mkApp (Lazy.force coq_heq_refl, [| eqt; c |]) in
if unif (push_rel_context ctx env) evars pt t then
@@ -3760,7 +3803,7 @@ let specialize_eqs id gl =
let acc, worked, ctx, ty = aux false [] (mkVar id) ty in
let ctx' = nf_rel_context_evar !evars ctx in
let ctx'' = List.map (function
- | LocalDef (n,k,t) when isEvar k -> LocalAssum (n,t)
+ | LocalDef (n,k,t) when isEvar !evars k -> LocalAssum (n,t)
| decl -> decl) ctx'
in
let ty' = it_mkProd_or_LetIn ty ctx'' in
@@ -3774,15 +3817,15 @@ let specialize_eqs id gl =
else tclFAIL 0 (str "Nothing to do in hypothesis " ++ pr_id id) gl
-let specialize_eqs id = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let specialize_eqs id = Proofview.Goal.enter { enter = begin fun gl ->
let msg = str "Specialization not allowed on dependent hypotheses" in
Proofview.tclOR (clear [id])
(fun _ -> Tacticals.New.tclZEROMSG msg) >>= fun () ->
Proofview.V82.tactic (specialize_eqs id)
end }
-let occur_rel n c =
- let res = not (noccurn n c) in
+let occur_rel sigma n c =
+ let res = not (noccurn sigma n c) in
res
(* This function splits the products of the induction scheme [elimt] into four
@@ -3793,20 +3836,20 @@ let occur_rel n c =
if there is no branch, we try to fill in acc3 with args/indargs.
We also return the conclusion.
*)
-let decompose_paramspred_branch_args elimt =
+let decompose_paramspred_branch_args sigma elimt =
let open Context.Rel.Declaration in
let rec cut_noccur elimt acc2 =
- match kind_of_term elimt with
+ match EConstr.kind sigma elimt with
| Prod(nme,tpe,elimt') ->
- let hd_tpe,_ = decompose_app ((strip_prod_assum tpe)) in
- if not (occur_rel 1 elimt') && isRel hd_tpe
+ let hd_tpe,_ = decompose_app sigma (snd (decompose_prod_assum sigma tpe)) in
+ if not (occur_rel sigma 1 elimt') && isRel sigma hd_tpe
then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2)
- else let acc3,ccl = decompose_prod_assum elimt in acc2 , acc3 , ccl
+ else let acc3,ccl = decompose_prod_assum sigma elimt in acc2 , acc3 , ccl
| App(_, _) | Rel _ -> acc2 , [] , elimt
| _ -> error_ind_scheme "" in
let rec cut_occur elimt acc1 =
- match kind_of_term elimt with
- | Prod(nme,tpe,c) when occur_rel 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1)
+ match EConstr.kind sigma elimt with
+ | Prod(nme,tpe,c) when occur_rel sigma 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1)
| Prod(nme,tpe,c) -> let acc2,acc3,ccl = cut_noccur elimt [] in acc1,acc2,acc3,ccl
| App(_, _) | Rel _ -> acc1,[],[],elimt
| _ -> error_ind_scheme "" in
@@ -3819,17 +3862,17 @@ let decompose_paramspred_branch_args elimt =
args. We suppose there is only one predicate here. *)
match acc2 with
| [] ->
- let hyps,ccl = decompose_prod_assum elimt in
- let hd_ccl_pred,_ = decompose_app ccl in
- begin match kind_of_term hd_ccl_pred with
+ let hyps,ccl = decompose_prod_assum sigma elimt in
+ let hd_ccl_pred,_ = decompose_app sigma ccl in
+ begin match EConstr.kind sigma hd_ccl_pred with
| Rel i -> let acc3,acc1 = List.chop (i-1) hyps in acc1 , [] , acc3 , ccl
| _ -> error_ind_scheme ""
end
| _ -> acc1, acc2 , acc3, ccl
-let exchange_hd_app subst_hd t =
- let hd,args= decompose_app t in mkApp (subst_hd,Array.of_list args)
+let exchange_hd_app sigma subst_hd t =
+ let hd,args= decompose_app sigma t in mkApp (subst_hd,Array.of_list args)
(* Builds an elim_scheme from its type and calling form (const+binding). We
first separate branches. We obtain branches, hyps before (params + preds),
@@ -3847,14 +3890,14 @@ let exchange_hd_app subst_hd t =
predicates are cited in the conclusion.
- finish to fill in the elim_scheme: indarg/farg/args and finally indref. *)
-let compute_elim_sig ?elimc elimt =
+let compute_elim_sig sigma ?elimc elimt =
let open Context.Rel.Declaration in
let params_preds,branches,args_indargs,conclusion =
- decompose_paramspred_branch_args elimt in
+ decompose_paramspred_branch_args sigma elimt in
- let ccl = exchange_hd_app (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in
+ let ccl = exchange_hd_app sigma (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in
let concl_with_args = it_mkProd_or_LetIn ccl args_indargs in
- let nparams = Int.Set.cardinal (free_rels concl_with_args) in
+ let nparams = Int.Set.cardinal (free_rels sigma concl_with_args) in
let preds,params = List.chop (List.length params_preds - nparams) params_preds in
(* A first approximation, further analysis will tweak it *)
@@ -3863,7 +3906,7 @@ let compute_elim_sig ?elimc elimt =
elimc = elimc; elimt = elimt; concl = conclusion;
predicates = preds; npredicates = List.length preds;
branches = branches; nbranches = List.length branches;
- farg_in_concl = isApp ccl && isApp (last_arg ccl);
+ farg_in_concl = isApp sigma ccl && isApp sigma (last_arg sigma ccl);
params = params; nparams = nparams;
(* all other fields are unsure at this point. Including these:*)
args = args_indargs; nargs = List.length args_indargs; } in
@@ -3884,9 +3927,9 @@ let compute_elim_sig ?elimc elimt =
match List.hd args_indargs with
| LocalDef (hiname,_,hi) -> error_ind_scheme ""
| LocalAssum (hiname,hi) ->
- let hi_ind, hi_args = decompose_app hi in
+ let hi_ind, hi_args = decompose_app sigma hi in
let hi_is_ind = (* hi est d'un type globalisable *)
- match kind_of_term hi_ind with
+ match EConstr.kind sigma hi_ind with
| Ind (mind,_) -> true
| Var _ -> true
| Const _ -> true
@@ -3899,7 +3942,7 @@ let compute_elim_sig ?elimc elimt =
else (* Last arg is the indarg *)
res := {!res with
indarg = Some (List.hd !res.args);
- indarg_in_concl = occur_rel 1 ccl;
+ indarg_in_concl = occur_rel sigma 1 ccl;
args = List.tl !res.args; nargs = !res.nargs - 1;
};
raise Exit);
@@ -3909,49 +3952,49 @@ let compute_elim_sig ?elimc elimt =
| None -> !res (* No indref *)
| Some (LocalDef _) -> error_ind_scheme ""
| Some (LocalAssum (_,ind)) ->
- let indhd,indargs = decompose_app ind in
- try {!res with indref = Some (global_of_constr indhd) }
+ let indhd,indargs = decompose_app sigma ind in
+ try {!res with indref = Some (fst (Termops.global_of_constr sigma indhd)) }
with e when CErrors.noncritical e ->
error "Cannot find the inductive type of the inductive scheme."
-let compute_scheme_signature scheme names_info ind_type_guess =
+let compute_scheme_signature evd scheme names_info ind_type_guess =
let open Context.Rel.Declaration in
- let f,l = decompose_app scheme.concl in
+ let f,l = decompose_app evd scheme.concl in
(* Vérifier que les arguments de Qi sont bien les xi. *)
let cond, check_concl =
match scheme.indarg with
| Some (LocalDef _) ->
error "Strange letin, cannot recognize an induction scheme."
| None -> (* Non standard scheme *)
- let cond hd = Term.eq_constr hd ind_type_guess && not scheme.farg_in_concl
+ let cond hd = EConstr.eq_constr evd hd ind_type_guess && not scheme.farg_in_concl
in (cond, fun _ _ -> ())
| Some (LocalAssum (_,ind)) -> (* Standard scheme from an inductive type *)
- let indhd,indargs = decompose_app ind in
- let cond hd = Term.eq_constr hd indhd in
+ let indhd,indargs = decompose_app evd ind in
+ let cond hd = EConstr.eq_constr evd hd indhd in
let check_concl is_pred p =
(* Check again conclusion *)
let ccl_arg_ok = is_pred (p + scheme.nargs + 1) f == IndArg in
let ind_is_ok =
- List.equal Term.eq_constr
+ List.equal (fun c1 c2 -> EConstr.eq_constr evd c1 c2)
(List.lastn scheme.nargs indargs)
- (Context.Rel.to_extended_list 0 scheme.args) in
+ (Context.Rel.to_extended_list mkRel 0 scheme.args) in
if not (ccl_arg_ok && ind_is_ok) then
error_ind_scheme "the conclusion of"
in (cond, check_concl)
in
let is_pred n c =
- let hd = fst (decompose_app c) in
- match kind_of_term hd with
+ let hd = fst (decompose_app evd c) in
+ match EConstr.kind evd hd with
| Rel q when n < q && q <= n+scheme.npredicates -> IndArg
| _ when cond hd -> RecArg
| _ -> OtherArg
in
let rec check_branch p c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Prod (_,t,c) ->
- (is_pred p t, true, dependent (mkRel 1) c) :: check_branch (p+1) c
+ (is_pred p t, true, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c
| LetIn (_,_,_,c) ->
- (OtherArg, false, dependent (mkRel 1) c) :: check_branch (p+1) c
+ (OtherArg, false, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c
| _ when is_pred p c == IndArg -> []
| _ -> raise Exit
in
@@ -3983,55 +4026,61 @@ let compute_scheme_signature scheme names_info ind_type_guess =
the non standard case, naming of generated hypos is slightly
different. *)
let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info =
- let scheme = compute_elim_sig ~elimc:elimc elimt in
- evd, (compute_scheme_signature scheme names_info ind_type_guess, scheme)
+ let scheme = compute_elim_sig evd ~elimc:elimc elimt in
+ evd, (compute_scheme_signature evd scheme names_info ind_type_guess, scheme)
let guess_elim isrec dep s hyp0 gl =
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in
- let mind,_ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in
+ let (mind, u), _ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in
let evd, elimc =
- if isrec && not (is_nonrec (fst mind)) then find_ind_eliminator (fst mind) s gl
+ if isrec && not (is_nonrec mind) then find_ind_eliminator mind s gl
else
let env = Tacmach.New.pf_env gl in
let sigma = Sigma.Unsafe.of_evar_map (Tacmach.New.project gl) in
+ let u = EInstance.kind (Tacmach.New.project gl) u in
if use_dependent_propositions_elimination () && dep
then
- let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma mind true s in
+ let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma (mind, u) true s in
+ let ind = EConstr.of_constr ind in
(Sigma.to_evar_map sigma, ind)
else
- let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma mind s in
+ let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma (mind, u) s in
+ let ind = EConstr.of_constr ind in
(Sigma.to_evar_map sigma, ind)
in
let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in
- evd, ((elimc, NoBindings), elimt), mkIndU mind
+ evd, ((elimc, NoBindings), elimt), mkIndU (mind, u)
let given_elim hyp0 (elimc,lbind as e) gl =
+ let sigma = Tacmach.New.project gl in
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in
- let ind_type_guess,_ = decompose_app ((strip_prod tmptyp0)) in
- Tacmach.New.project gl, (e, Tacmach.New.pf_unsafe_type_of gl elimc), ind_type_guess
+ let ind_type_guess,_ = decompose_app sigma (snd (decompose_prod sigma tmptyp0)) in
+ let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in
+ Tacmach.New.project gl, (e, elimt), ind_type_guess
type scheme_signature =
(Id.t list * (elim_arg_kind * bool * bool * Id.t) list) array
type eliminator_source =
- | ElimUsing of (eliminator * types) * scheme_signature
+ | ElimUsing of (eliminator * EConstr.types) * scheme_signature
| ElimOver of bool * Id.t
let find_induction_type isrec elim hyp0 gl =
+ let sigma = Tacmach.New.project gl in
let scheme,elim =
match elim with
| None ->
let sort = Tacticals.New.elimination_sort_of_goal gl in
let _, (elimc,elimt),_ =
guess_elim isrec (* dummy: *) true sort hyp0 gl in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
(* We drop the scheme waiting to know if it is dependent *)
scheme, ElimOver (isrec,hyp0)
| Some e ->
let evd, (elimc,elimt),ind_guess = given_elim hyp0 e gl in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
if Option.is_empty scheme.indarg then error "Cannot find induction type";
- let indsign = compute_scheme_signature scheme hyp0 ind_guess in
+ let indsign = compute_scheme_signature evd scheme hyp0 ind_guess in
let elim = ({elimindex = Some(-1); elimbody = elimc; elimrename = None},elimt) in
scheme, ElimUsing (elim,indsign)
in
@@ -4043,7 +4092,8 @@ let get_elim_signature elim hyp0 gl =
compute_elim_signature (given_elim hyp0 elim gl) hyp0
let is_functional_induction elimc gl =
- let scheme = compute_elim_sig ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in
+ let sigma = Tacmach.New.project gl in
+ let scheme = compute_elim_sig sigma ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in
(* The test is not safe: with non-functional induction on non-standard
induction scheme, this may fail *)
Option.is_empty scheme.indarg
@@ -4058,7 +4108,7 @@ let get_eliminator elim dep s gl =
| ElimOver (isrec,id) ->
let evd, (elimc,elimt),_ as elims = guess_elim isrec dep s id gl in
let _, (l, s) = compute_elim_signature elims id in
- let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (RelDecl.get_type d)))
+ let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (Tacmach.New.project gl) (RelDecl.get_type d)))
(List.rev s.branches)
in
evd, isrec, ({elimindex = None; elimbody = elimc; elimrename = Some (isrec,Array.of_list branchlengthes)}, elimt), l
@@ -4067,11 +4117,11 @@ let get_eliminator elim dep s gl =
of lid are parameters (first ones), the other are
arguments. Returns the clause obtained. *)
let recolle_clenv i params args elimclause gl =
- let _,arr = destApp elimclause.templval.rebus in
+ let _,arr = destApp elimclause.evd elimclause.templval.rebus in
let lindmv =
Array.map
(fun x ->
- match kind_of_term x with
+ match EConstr.kind elimclause.evd x with
| Meta mv -> mv
| _ -> user_err ~hdr:"elimination_clause"
(str "The type of the elimination clause is not well-formed."))
@@ -4092,7 +4142,7 @@ let recolle_clenv i params args elimclause gl =
(* from_n (Some 0) means that x should be taken "as is" without
trying to unify (which would lead to trying to apply it to
evars if y is a product). *)
- let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from_n gl (Some 0) (x,y)) gl in
+ let indclause = mk_clenv_from_n gl (Some 0) (x,y) in
let elimclause' = clenv_fchain ~with_univs:false i acc indclause in
elimclause')
(List.rev clauses)
@@ -4103,17 +4153,18 @@ let recolle_clenv i params args elimclause gl =
produce new ones). Then refine with the resulting term with holes.
*)
let induction_tac with_evars params indvars elim =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let ({elimindex=i;elimbody=(elimc,lbindelimc);elimrename=rename},elimt) = elim in
- let i = match i with None -> index_of_ind_arg elimt | Some i -> i in
+ let i = match i with None -> index_of_ind_arg sigma elimt | Some i -> i in
(* elimclause contains this: (elimc ?i ?j ?k...?l) *)
- let elimc = contract_letin_in_lam_header elimc in
+ let elimc = contract_letin_in_lam_header sigma elimc in
let elimc = mkCast (elimc, DEFAULTcast, elimt) in
- let elimclause = pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in
+ let elimclause = Tacmach.New.pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in
(* elimclause' is built from elimclause by instanciating all args and params. *)
let elimclause' = recolle_clenv i params indvars elimclause gl in
(* one last resolution (useless?) *)
- let resolved = Tacmach.New.of_old (clenv_unique_resolver ~flags:(elim_flags ()) elimclause') gl in
+ let resolved = clenv_unique_resolver ~flags:(elim_flags ()) elimclause' gl in
enforce_prop_bound_names rename (Clenvtac.clenv_refine with_evars resolved)
end }
@@ -4126,9 +4177,9 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let sigma = Sigma.to_evar_map sigma in
- let concl = Tacmach.New.pf_nf_concl gl in
- let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env in
- let dep_in_concl = Option.cata (fun id -> occur_var env id concl) false hyp0 in
+ let concl = Tacmach.New.pf_concl gl in
+ let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in
+ let dep_in_concl = Option.cata (fun id -> occur_var env sigma id concl) false hyp0 in
let dep = dep_in_hyps || dep_in_concl in
let tmpcl = it_mkNamedProd_or_LetIn concl deps in
let s = Retyping.get_sort_family_of env sigma tmpcl in
@@ -4180,7 +4231,7 @@ let msg_not_right_number_induction_arguments scheme =
must be given, so we help a bit the unifier by making the "pattern"
by hand before calling induction_tac *)
let induction_without_atomization isrec with_evars elim names lid =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma, (indsign,scheme) = get_elim_signature elim (List.hd lid) gl in
let nargs_indarg_farg =
scheme.nargs + (if scheme.farg_in_concl then 1 else 0) in
@@ -4215,8 +4266,8 @@ let induction_without_atomization isrec with_evars elim names lid =
(* assume that no occurrences are selected *)
let clear_unselected_context id inhyps cls =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- if occur_var (Tacmach.New.pf_env gl) id (Tacmach.New.pf_concl gl) &&
+ Proofview.Goal.enter { enter = begin fun gl ->
+ if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (Tacmach.New.pf_concl gl) &&
cls.concl_occs == NoOccurrences
then user_err
(str "Conclusion must be mentioned: it depends on " ++ pr_id id
@@ -4228,7 +4279,7 @@ let clear_unselected_context id inhyps cls =
if Id.List.mem id' inhyps then (* if selected, do not erase *) None
else
(* erase if not selected and dependent on id or selected hyps *)
- let test id = occur_var_in_decl (Tacmach.New.pf_env gl) id d in
+ let test id = occur_var_in_decl (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id d in
if List.exists test (id::inhyps) then Some id' else None in
let ids = List.map_filter to_erase (Proofview.Goal.hyps gl) in
clear ids
@@ -4255,7 +4306,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ =
let rec find_clause typ =
try
let indclause = make_clenv_binding env sigma (c,typ) lbind in
- if must_be_closed && occur_meta (clenv_value indclause) then
+ if must_be_closed && occur_meta indclause.evd (clenv_value indclause) then
error "Need a fully applied argument.";
(* We lose the possibility of coercions in with-bindings *)
let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in
@@ -4273,7 +4324,7 @@ let check_expected_type env sigma (elimc,bl) elimt =
if n == 0 then error "Scheme cannot be applied.";
let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in
let sigma = solve_evar_clause env sigma true cl bl in
- let (_,u,_) = destProd cl.cl_concl in
+ let (_,u,_) = destProd sigma cl.cl_concl in
fun t -> Evarconv.e_cumul env (ref sigma) t u
let check_enough_applied env sigma elim =
@@ -4283,10 +4334,10 @@ let check_enough_applied env sigma elim =
| None ->
(* No eliminator given *)
fun u ->
- let t,_ = decompose_app (whd_all env sigma u) in isInd t
+ let t,_ = decompose_app sigma (whd_all env sigma u) in isInd sigma t
| Some elimc ->
let elimt = Retyping.get_type_of env sigma (fst elimc) in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
match scheme.indref with
| None ->
(* in the absence of information, do not assume it may be
@@ -4305,7 +4356,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
- let ccl = Proofview.Goal.raw_concl gl in
+ let ccl = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let check = check_enough_applied env sigma elim in
let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in
@@ -4336,7 +4387,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
end };
if with_evars then Proofview.shelve_unifiable else guard_no_unifiable;
if is_arg_pure_hyp
- then Tacticals.New.tclTRY (clear [destVar c0])
+ then Proofview.tclEVARMAP >>= fun sigma -> Tacticals.New.tclTRY (clear [destVar sigma c0])
else Proofview.tclUNIT ();
if isrec then Proofview.cycle (-1) else Proofview.tclUNIT ()
])
@@ -4360,10 +4411,10 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
Sigma (tac, sigma', p +> q)
end }
-let has_generic_occurrences_but_goal cls id env ccl =
+let has_generic_occurrences_but_goal cls id env sigma ccl =
clause_with_generic_context_selection cls &&
(* TODO: whd_evar of goal *)
- (cls.concl_occs != NoOccurrences || not (occur_var env id ccl))
+ (cls.concl_occs != NoOccurrences || not (occur_var env sigma id ccl))
let induction_gen clear_flag isrec with_evars elim
((_pending,(c,lbind)),(eqname,names) as arg) cls =
@@ -4373,14 +4424,15 @@ let induction_gen clear_flag isrec with_evars elim
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
- let ccl = Proofview.Goal.raw_concl gl in
+ let evd = Sigma.to_evar_map sigma in
+ let ccl = Proofview.Goal.concl gl in
let cls = Option.default allHypsAndConcl cls in
let t = typ_of env sigma c in
let is_arg_pure_hyp =
- isVar c && not (mem_named_context_val (destVar c) (Global.named_context_val ()))
+ isVar evd c && not (mem_named_context_val (destVar evd c) (Global.named_context_val ()))
&& lbind == NoBindings && not with_evars && Option.is_empty eqname
&& clear_flag == None
- && has_generic_occurrences_but_goal cls (destVar c) env ccl in
+ && has_generic_occurrences_but_goal cls (destVar evd c) env evd ccl in
let enough_applied = check_enough_applied env sigma elim t in
if is_arg_pure_hyp && enough_applied then
(* First case: induction on a variable already in an inductive type and
@@ -4388,7 +4440,7 @@ let induction_gen clear_flag isrec with_evars elim
This is a situation where the induction argument is a
clearable variable of the goal w/o occurrence selection
and w/o equality kept: no need to generalize *)
- let id = destVar c in
+ let id = destVar evd c in
Tacticals.New.tclTHEN
(clear_unselected_context id inhyps cls)
(induction_with_atomization_of_ind_arg
@@ -4398,7 +4450,7 @@ let induction_gen clear_flag isrec with_evars elim
declaring the induction argument as a new local variable *)
let id =
(* Type not the right one if partially applied but anyway for internal use*)
- let x = id_of_name_using_hdchar (Global.env()) t Anonymous in
+ let x = id_of_name_using_hdchar (Global.env()) evd t Anonymous in
new_fresh_id [] x gl in
let info_arg = (is_arg_pure_hyp, not enough_applied) in
pose_induction_arg_then
@@ -4423,7 +4475,8 @@ let induction_gen_l isrec with_evars elim names lc =
match l with
| [] -> Proofview.tclUNIT ()
| c::l' ->
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when not (mem_named_context_val id (Global.named_context_val ()))
&& not with_evars ->
let _ = newlc:= id::!newlc in
@@ -4432,11 +4485,12 @@ let induction_gen_l isrec with_evars elim names lc =
| _ ->
Proofview.Goal.enter { enter = begin fun gl ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
+ let sigma = Tacmach.New.project gl in
let x =
- id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in
+ id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in
let id = new_fresh_id [] x gl in
- let newl' = List.map (replace_term c (mkVar id)) l' in
+ let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in
let _ = newlc:=id::!newlc in
Tacticals.New.tclTHEN
(letin_tac None (Name id) c None allHypsAndConcl)
@@ -4458,7 +4512,7 @@ let induction_destruct isrec with_evars (lc,elim) =
match lc with
| [] -> assert false (* ensured by syntax, but if called inside caml? *)
| [c,(eqname,names as allnames),cls] ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
match elim with
@@ -4559,9 +4613,9 @@ let simple_destruct = function
*)
let elim_scheme_type elim t =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let clause = Tacmach.New.of_old (fun gl -> mk_clenv_type_of gl elim) gl in
- match kind_of_term (last_arg clause.templval.rebus) with
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let clause = mk_clenv_type_of gl elim in
+ match EConstr.kind clause.evd (last_arg clause.evd clause.templval.rebus) with
| Meta mv ->
let clause' =
(* t is inductive, then CUMUL or CONV is irrelevant *)
@@ -4582,9 +4636,11 @@ let case_type t =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Tacmach.New.pf_env gl in
- let (ind,t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in
+ let ((ind, u), t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in
+ let u = EInstance.kind (Sigma.to_evar_map sigma) u in
let s = Tacticals.New.elimination_sort_of_goal gl in
- let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma ind s in
+ let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma (ind, u) s in
+ let elimc = EConstr.of_constr elimc in
Sigma (elim_scheme_type elimc t, evd, p)
end }
@@ -4598,7 +4654,7 @@ let case_type t =
let (forward_setoid_reflexivity, setoid_reflexivity) = Hook.make ()
let maybe_betadeltaiota_concl allowred gl =
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
let sigma = Tacmach.New.project gl in
if not allowred then concl
else
@@ -4610,8 +4666,9 @@ let reflexivity_red allowred =
(* PL: usual reflexivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match match_with_equality_type concl with
+ match match_with_equality_type sigma concl with
| None -> Proofview.tclZERO NoEquationFound
| Some _ -> one_constructor 1 NoBindings
end }
@@ -4648,9 +4705,9 @@ let prove_symmetry hdcncl eq_kind =
Tacticals.New.onLastHyp simplest_case;
one_constructor 1 NoBindings ])
-let match_with_equation c =
+let match_with_equation sigma c =
try
- let res = match_with_equation c in
+ let res = match_with_equation sigma c in
Proofview.tclUNIT res
with NoEquationFound ->
Proofview.tclZERO NoEquationFound
@@ -4660,8 +4717,9 @@ let symmetry_red allowred =
(* PL: usual symmetry don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match_with_equation concl >>= fun with_eqn ->
+ match_with_equation sigma concl >>= fun with_eqn ->
match with_eqn with
| Some eq_data,_,_ ->
Tacticals.New.tclTHEN
@@ -4683,16 +4741,18 @@ let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make ()
let symmetry_in id =
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in
- let sign,t = decompose_prod_assum ctype in
+ let sign,t = decompose_prod_assum sigma ctype in
Proofview.tclORELSE
begin
- match_with_equation t >>= fun (_,hdcncl,eq) ->
- let symccl = match eq with
+ match_with_equation sigma t >>= fun (_,hdcncl,eq) ->
+ let symccl =
+ match eq with
| MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |])
| PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |])
| HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in
- Tacticals.New.tclTHENS (cut (it_mkProd_or_LetIn symccl sign))
+ Tacticals.New.tclTHENS (cut (EConstr.it_mkProd_or_LetIn symccl sign))
[ intro_replacing id;
Tacticals.New.tclTHENLIST [ intros; symmetry; apply (mkVar id); assumption ] ]
end
@@ -4752,8 +4812,9 @@ let transitivity_red allowred t =
(* PL: usual transitivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match_with_equation concl >>= fun with_eqn ->
+ match_with_equation sigma concl >>= fun with_eqn ->
match with_eqn with
| Some eq_data,_,_ ->
Tacticals.New.tclTHEN
@@ -4787,6 +4848,12 @@ let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n)
(** d1 is the section variable in the global context, d2 in the goal context *)
let interpretable_as_section_decl evd d1 d2 =
let open Context.Named.Declaration in
+ let e_eq_constr_univs sigma c1 c2 = match eq_constr_universes !sigma c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !sigma cstr); true
+ with UniversesDiffer -> false
+ in
match d2, d1 with
| LocalDef _, LocalAssum _ -> false
| LocalDef (_,b1,t1), LocalDef (_,b2,t2) ->
@@ -4804,6 +4871,8 @@ let rec decompose len c t accu =
| _ -> assert false
let rec shrink ctx sign c t accu =
+ let open Term in
+ let open CVars in
match ctx, sign with
| [], [] -> (c, t, accu)
| p :: ctx, decl :: sign ->
@@ -4842,7 +4911,7 @@ let abstract_subproof id gk tac =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let current_sign = Global.named_context_val ()
and global_sign = Proofview.Goal.hyps gl in
@@ -4870,6 +4939,7 @@ let abstract_subproof id gk tac =
let ctx = Evd.universe_context_set evd in
evd, ctx, nf concl
in
+ let concl = EConstr.of_constr concl in
let solve_tac = tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac) in
let ectx = Evd.evar_universe_context evd in
let (const, safe, ectx) =
@@ -4884,8 +4954,9 @@ let abstract_subproof id gk tac =
in
let const, args =
if !shrink_abstract then shrink_entry sign const
- else (const, List.rev (Context.Named.to_instance sign))
+ else (const, List.rev (Context.Named.to_instance Constr.mkVar sign))
in
+ let args = List.map EConstr.of_constr args in
let cd = Entries.DefinitionEntry const in
let decl = (cd, IsProof Lemma) in
let cst () =
@@ -4897,6 +4968,7 @@ let abstract_subproof id gk tac =
let cst = Impargs.with_implicit_protection cst () in
(* let evd, lem = Evd.fresh_global (Global.env ()) evd (ConstRef cst) in *)
let lem, ctx = Universes.unsafe_constr_of_global (ConstRef cst) in
+ let lem = EConstr.of_constr lem in
let evd = Evd.set_universe_context evd ectx in
let open Safe_typing in
let eff = private_con_of_con (Global.safe_env ()) cst in
@@ -4928,7 +5000,7 @@ let tclABSTRACT name_op tac =
abstract_subproof s gk tac
let unify ?(state=full_transparent_state) x y =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
try
let core_flags =
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 354ac50b4..ba4a9706d 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -9,6 +9,7 @@
open Loc
open Names
open Term
+open EConstr
open Environ
open Proof_type
open Evd
@@ -28,15 +29,15 @@ open Locus
(** {6 General functions. } *)
-val is_quantified_hypothesis : Id.t -> ([`NF],'b) Proofview.Goal.t -> bool
+val is_quantified_hypothesis : Id.t -> ('a, 'r) Proofview.Goal.t -> bool
(** {6 Primitive tactics. } *)
val introduction : ?check:bool -> Id.t -> unit Proofview.tactic
val convert_concl : ?check:bool -> types -> cast_kind -> unit Proofview.tactic
-val convert_hyp : ?check:bool -> Context.Named.Declaration.t -> unit Proofview.tactic
+val convert_hyp : ?check:bool -> named_declaration -> unit Proofview.tactic
val convert_concl_no_check : types -> cast_kind -> unit Proofview.tactic
-val convert_hyp_no_check : Context.Named.Declaration.t -> unit Proofview.tactic
+val convert_hyp_no_check : named_declaration -> unit Proofview.tactic
val mutual_fix :
Id.t -> int -> (Id.t * int * constr) list -> int -> unit Proofview.tactic
val fix : Id.t option -> int -> unit Proofview.tactic
@@ -50,7 +51,7 @@ val convert_leq : constr -> constr -> unit Proofview.tactic
val fresh_id_in_env : Id.t list -> Id.t -> env -> Id.t
val fresh_id : Id.t list -> Id.t -> goal sigma -> Id.t
-val find_intro_names : Context.Rel.t -> goal sigma -> Id.t list
+val find_intro_names : rel_context -> goal sigma -> Id.t list
val intro : unit Proofview.tactic
val introf : unit Proofview.tactic
@@ -74,7 +75,7 @@ val intros : unit Proofview.tactic
(** [depth_of_quantified_hypothesis b h g] returns the index of [h] in
the conclusion of goal [g], up to head-reduction if [b] is [true] *)
val depth_of_quantified_hypothesis :
- bool -> quantified_hypothesis -> ([`NF],'b) Proofview.Goal.t -> int
+ bool -> quantified_hypothesis -> ('a, 'r) Proofview.Goal.t -> int
val intros_until : quantified_hypothesis -> unit Proofview.tactic
@@ -184,7 +185,7 @@ val revert : Id.t list -> unit Proofview.tactic
(** {6 Resolution tactics. } *)
val apply_type : constr -> constr list -> unit Proofview.tactic
-val bring_hyps : Context.Named.t -> unit Proofview.tactic
+val bring_hyps : named_context -> unit Proofview.tactic
val apply : constr -> unit Proofview.tactic
val eapply : constr -> unit Proofview.tactic
@@ -243,15 +244,15 @@ type elim_scheme = {
elimc: constr with_bindings option;
elimt: types;
indref: global_reference option;
- params: Context.Rel.t; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
+ params: rel_context; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
nparams: int; (** number of parameters *)
- predicates: Context.Rel.t; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
+ predicates: rel_context; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
npredicates: int; (** Number of predicates *)
- branches: Context.Rel.t; (** branchr,...,branch1 *)
+ branches: rel_context; (** branchr,...,branch1 *)
nbranches: int; (** Number of branches *)
- args: Context.Rel.t; (** (xni, Ti_ni) ... (x1, Ti_1) *)
+ args: rel_context; (** (xni, Ti_ni) ... (x1, Ti_1) *)
nargs: int; (** number of arguments *)
- indarg: Context.Rel.Declaration.t option; (** Some (H,I prm1..prmp x1...xni)
+ indarg: rel_declaration option; (** Some (H,I prm1..prmp x1...xni)
if HI is in premisses, None otherwise *)
concl: types; (** Qi x1...xni HI (f...), HI and (f...)
are optional and mutually exclusive *)
@@ -259,7 +260,7 @@ type elim_scheme = {
farg_in_concl: bool; (** true if (f...) appears at the end of conclusion *)
}
-val compute_elim_sig : ?elimc: constr with_bindings -> types -> elim_scheme
+val compute_elim_sig : evar_map -> ?elimc:constr with_bindings -> types -> elim_scheme
(** elim principle with the index of its inductive arg *)
type eliminator = {
diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml
index e4b45489d..2c863c42a 100644
--- a/tactics/term_dnet.ml
+++ b/tactics/term_dnet.ml
@@ -344,18 +344,19 @@ struct
) (pr_dconstr pr_term_pattern) p*)
let search_pat cpat dpat dn =
- let whole_c = cpat in
+ let whole_c = EConstr.of_constr cpat in
(* if we are at the root, add an empty context *)
let dpat = under_prod (empty_ctx dpat) in
TDnet.Idset.fold
(fun id acc ->
let c_id = Opt.reduce (Ident.constr_of id) in
+ let c_id = EConstr.of_constr c_id in
let (ctx,wc) =
- try Termops.align_prod_letin whole_c c_id
+ try Termops.align_prod_letin Evd.empty whole_c c_id (** FIXME *)
with Invalid_argument _ -> [],c_id in
let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in
try
- let _ = Termops.filtering ctx Reduction.CUMUL wc whole_c in
+ let _ = Termops.filtering Evd.empty ctx Reduction.CUMUL wc whole_c in
id :: acc
with Termops.CannotFilter -> (* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *) acc
) (TDnet.find_match dpat dn) []
diff --git a/test-suite/output/inference.out b/test-suite/output/inference.out
index e83e7176d..c70467912 100644
--- a/test-suite/output/inference.out
+++ b/test-suite/output/inference.out
@@ -6,13 +6,13 @@ fun e : option L => match e with
: option L -> option L
fun (m n p : nat) (H : S m <= S n + p) => le_S_n m (n + p) H
: forall m n p : nat, S m <= S n + p -> m <= n + p
-fun n : nat => let x : T n := A n in ?y ?y0 : T n
+fun n : nat => let x : T n := A n in ?t ?y : T n
: forall n : nat, T n
where
-?y : [n : nat x := A n : T n |- ?T -> T n]
-?y0 : [n : nat x := A n : T n |- ?T]
-fun n : nat => ?y ?y0 : T n
+?t : [n : nat x := A n : T n |- ?T -> T n]
+?y : [n : nat x := A n : T n |- ?T]
+fun n : nat => ?t ?y : T n
: forall n : nat, T n
where
-?y : [n : nat |- ?T -> T n]
-?y0 : [n : nat |- ?T]
+?t : [n : nat |- ?T -> T n]
+?y : [n : nat |- ?T]
diff --git a/test-suite/success/change_pattern.v b/test-suite/success/change_pattern.v
new file mode 100644
index 000000000..874abf49f
--- /dev/null
+++ b/test-suite/success/change_pattern.v
@@ -0,0 +1,34 @@
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Axiom vector : Type -> nat -> Type.
+
+Record KleeneStore i j a := kleeneStore
+ { dim : nat
+ ; peek : vector j dim -> a
+ ; pos : vector i dim
+ }.
+
+Definition KSmap i j a b (f : a -> b) (s : KleeneStore i j a) : KleeneStore i j b :=
+ kleeneStore (fun v => f (peek v)) (pos s).
+
+Record KleeneCoalg (i o : Type -> Type) := kleeneCoalg
+ { coalg :> forall a b, (o a) -> KleeneStore (i a) (i b) (o b) }.
+
+Axiom free_b_dim : forall i o (k : KleeneCoalg i o) a b b' (x : o a), dim (coalg k b x) = dim (coalg k b' x).
+Axiom t : Type -> Type.
+Axiom traverse : KleeneCoalg (fun x => x) t.
+
+Definition size a (x:t a) : nat := dim (traverse a a x).
+
+Lemma iso1_iso2_2 a (y : {x : t unit & vector a (size x)}) : False.
+Proof.
+destruct y.
+pose (X := KSmap (traverse a unit) (traverse unit a x)).
+set (e :=(eq_sym (free_b_dim traverse (a:=unit) a unit x))).
+clearbody e.
+(** The pattern generated by change must have holes where there were implicit
+ arguments in the original user-provided term. This particular example fails
+ if this is not the case because the inferred argument does not coincide with
+ the one in the considered term. *)
+progress (change (dim (traverse unit a x)) with (dim X) in e).
diff --git a/test-suite/success/old_typeclass.v b/test-suite/success/old_typeclass.v
new file mode 100644
index 000000000..01e35810b
--- /dev/null
+++ b/test-suite/success/old_typeclass.v
@@ -0,0 +1,13 @@
+Require Import Setoid Coq.Classes.Morphisms.
+Set Typeclasses Legacy Resolution.
+
+Declare Instance and_Proper_eq: Proper (Logic.eq ==> Logic.eq ==> Logic.eq) and.
+
+Axiom In : Prop.
+Axiom union_spec : In <-> True.
+
+Lemma foo : In /\ True.
+Proof.
+progress rewrite union_spec.
+repeat constructor.
+Qed.
diff --git a/test-suite/success/rewrite_evar.v b/test-suite/success/rewrite_evar.v
new file mode 100644
index 000000000..f7ad261cb
--- /dev/null
+++ b/test-suite/success/rewrite_evar.v
@@ -0,0 +1,8 @@
+Require Import Coq.Setoids.Setoid.
+
+Goal forall (T2 MT1 MT2 : Type) (x : T2) (M2 m2 : MT2) (M1 m1 : MT1) (F : T2 -> MT1 -> MT2 -> Prop),
+ (forall (defaultB : T2) (m3 : MT1) (m4 : MT2), F defaultB m3 m4 <-> True) -> F x M1 M2 -> F x m1 m2.
+ intros ????????? H' H.
+ rewrite (H' _) in *.
+ (** The above rewrite should also rewrite in H. *)
+ Fail progress rewrite H' in H.
diff --git a/vernac/assumptions.ml b/vernac/assumptions.ml
index 8865cd646..deb2ed3e0 100644
--- a/vernac/assumptions.ml
+++ b/vernac/assumptions.ml
@@ -144,6 +144,27 @@ let label_of = function
| ConstructRef ((kn,_),_) -> pi3 (repr_mind kn)
| VarRef id -> Label.of_id id
+let fold_constr_with_full_binders g f n acc c =
+ let open Context.Rel.Declaration in
+ match kind_of_term 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
+ | Lambda (na,t,c) -> f (g (LocalAssum (na,t)) n) (f n acc t) c
+ | LetIn (na,b,t,c) -> f (g (LocalDef (na,b,t)) n) (f n (f n acc b) t) c
+ | App (c,l) -> Array.fold_left (f n) (f n acc c) l
+ | Proj (p,c) -> f n acc c
+ | Evar (_,l) -> Array.fold_left (f n) acc l
+ | Case (_,p,c,bl) -> Array.fold_left (f n) (f n (f n acc p) c) bl
+ | Fix (_,(lna,tl,bl)) ->
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ 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
+ | CoFix (_,(lna,tl,bl)) ->
+ let n' = CArray.fold_left2 (fun c n t -> g (LocalAssum (n,t)) c) n lna tl in
+ 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
| Var id ->
let body () = id |> Global.lookup_named |> NamedDecl.get_value in
@@ -166,10 +187,10 @@ let rec traverse current ctx accu t = match kind_of_term t with
traverse_object
~inhabits:(current,ctx,Vars.subst1 mkProp oty) accu body (ConstRef kn)
| _ ->
- Termops.fold_constr_with_full_binders
+ fold_constr_with_full_binders
Context.Rel.add (traverse current) ctx accu t
end
-| _ -> Termops.fold_constr_with_full_binders
+| _ -> fold_constr_with_full_binders
Context.Rel.add (traverse current) ctx accu t
and traverse_object ?inhabits (curr, data, ax2ty) body obj =
diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml
index 6d71601cc..b9c4c6cc5 100644
--- a/vernac/auto_ind_decl.ml
+++ b/vernac/auto_ind_decl.ml
@@ -105,7 +105,7 @@ let mkFullInd (ind,u) n =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
if nparrec > 0
then mkApp (mkIndU (ind,u),
- Array.of_list(Context.Rel.to_extended_list (nparrec+n) lnamesparrec))
+ Array.of_list(Context.Rel.to_extended_list mkRel (nparrec+n) lnamesparrec))
else mkIndU (ind,u)
let check_bool_is_defined () =
@@ -140,7 +140,7 @@ let build_beq_scheme mode kn =
| Name s -> Id.of_string ("eq_"^(Id.to_string s))
| Anonymous -> Id.of_string "eq_A"
in
- let ext_rel_list = Context.Rel.to_extended_list 0 lnamesparrec in
+ let ext_rel_list = Context.Rel.to_extended_list mkRel 0 lnamesparrec in
let lift_cnt = ref 0 in
let eqs_typ = List.map (fun aa ->
let a = lift !lift_cnt aa in
@@ -179,9 +179,10 @@ let build_beq_scheme mode kn =
*)
let compute_A_equality rel_list nlist eqA ndx t =
let lifti = ndx in
+ let sigma = Evd.empty (** FIXME *) in
let rec aux c =
let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Rel x -> mkRel (x-nlist+ndx), Safe_typing.empty_private_constants
| Var x ->
let eid = id_of_string ("eq_"^(string_of_id x)) in
@@ -190,7 +191,7 @@ let build_beq_scheme mode kn =
with Not_found -> raise (ParameterWithoutEquality (VarRef x))
in
mkVar eid, Safe_typing.empty_private_constants
- | Cast (x,_,_) -> aux (applist (x,a))
+ | Cast (x,_,_) -> aux (EConstr.applist (x,a))
| App _ -> assert false
| Ind ((kn',i as ind'),u) (*FIXME: universes *) ->
if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Safe_typing.empty_private_constants
@@ -206,7 +207,7 @@ let build_beq_scheme mode kn =
in
let args =
Array.append
- (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in
+ (Array.of_list (List.map (fun x -> lift lifti (EConstr.Unsafe.to_constr x)) a)) eqa in
if Int.equal (Array.length args) 0 then eq, eff
else mkApp (eq, args), eff
with Not_found -> raise(EqNotFound (ind', fst ind))
@@ -215,10 +216,11 @@ let build_beq_scheme mode kn =
| Prod _ -> raise InductiveWithProduct
| Lambda _-> raise (EqUnknown "abstraction")
| LetIn _ -> raise (EqUnknown "let-in")
- | Const kn ->
- (match Environ.constant_opt_value_in env kn with
- | None -> raise (ParameterWithoutEquality (ConstRef (fst kn)))
- | Some c -> aux (applist (c,a)))
+ | Const (kn, u) ->
+ let u = EConstr.EInstance.kind sigma u in
+ (match Environ.constant_opt_value_in env (kn, u) with
+ | None -> raise (ParameterWithoutEquality (ConstRef kn))
+ | Some c -> aux (EConstr.applist (EConstr.of_constr c,a)))
| Proj _ -> raise (EqUnknown "projection")
| Construct _ -> raise (EqUnknown "constructor")
| Case _ -> raise (EqUnknown "match")
@@ -242,7 +244,7 @@ let build_beq_scheme mode kn =
Cn => match Y with ... end |] part *)
let ci = make_case_info env (fst ind) MatchStyle in
let constrs n = get_constructors env (make_ind_family (ind,
- Context.Rel.to_extended_list (n+nb_ind-1) mib.mind_params_ctxt)) in
+ Context.Rel.to_extended_list mkRel (n+nb_ind-1) mib.mind_params_ctxt)) in
let constrsi = constrs (3+nparrec) in
let n = Array.length constrsi in
let ar = Array.make n (Lazy.force ff) in
@@ -262,7 +264,7 @@ let build_beq_scheme mode kn =
nparrec
(nparrec+3+2*nb_cstr_args)
(nb_cstr_args+ndx+1)
- cc
+ (EConstr.of_constr cc)
in
eff := Safe_typing.concat_private eff' !eff;
Array.set eqs ndx
@@ -324,11 +326,12 @@ let _ = beq_scheme_kind_aux := fun () -> beq_scheme_kind
(* This function tryies to get the [inductive] between a constr
the constr should be Ind i or App(Ind i,[|args|])
*)
-let destruct_ind c =
- try let u,v = destApp c in
- let indc = destInd u in
+let destruct_ind sigma c =
+ let open EConstr in
+ try let u,v = destApp sigma c in
+ let indc = destInd sigma u in
indc,v
- with DestKO -> let indc = destInd c in
+ with DestKO -> let indc = destInd sigma c in
indc,[||]
(*
@@ -341,11 +344,12 @@ so from Ai we can find the the correct eq_Ai bl_ai or lb_ai
*)
(* used in the leib -> bool side*)
let do_replace_lb mode lb_scheme_key aavoid narg p q =
+ let open EConstr in
let avoid = Array.of_list aavoid in
- let do_arg v offset =
+ let do_arg sigma v offset =
try
let x = narg*offset in
- let s = destVar v in
+ let s = destVar sigma v in
let n = Array.length avoid in
let rec find i =
if Id.equal avoid.(n-i) s then avoid.(n-i-x)
@@ -358,19 +362,20 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
(* if this happen then the args have to be already declared as a
Parameter*)
(
- let mp,dir,lbl = repr_con (fst (destConst v)) in
+ let mp,dir,lbl = repr_con (fst (destConst sigma v)) in
mkConst (make_con mp dir (mk_label (
if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
else ((Label.to_string lbl)^"_lb")
)))
)
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let type_of_pq = Tacmach.New.of_old (fun gl -> pf_unsafe_type_of gl p) gl in
- let u,v = destruct_ind type_of_pq
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let type_of_pq = Tacmach.New.pf_unsafe_type_of gl p in
+ let sigma = Tacmach.New.project gl in
+ let u,v = destruct_ind sigma type_of_pq
in let lb_type_of_p =
try
- let c, eff = find_scheme ~mode lb_scheme_key (out_punivs u) (*FIXME*) in
+ let c, eff = find_scheme ~mode lb_scheme_key (fst u) (*FIXME*) in
Proofview.tclUNIT (mkConst c, eff)
with Not_found ->
(* spiwack: the format of this error message should probably
@@ -379,17 +384,18 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
(str "Leibniz->boolean:" ++
str "You have to declare the" ++
str "decidability over " ++
- Printer.pr_constr type_of_pq ++
+ Printer.pr_econstr type_of_pq ++
str " first.")
in
Tacticals.New.tclZEROMSG err_msg
in
lb_type_of_p >>= fun (lb_type_of_p,eff) ->
+ Proofview.tclEVARMAP >>= fun sigma ->
let lb_args = Array.append (Array.append
(Array.map (fun x -> x) v)
- (Array.map (fun x -> do_arg x 1) v))
- (Array.map (fun x -> do_arg x 2) v)
- in let app = if Array.equal eq_constr lb_args [||]
+ (Array.map (fun x -> do_arg sigma x 1) v))
+ (Array.map (fun x -> do_arg sigma x 2) v)
+ in let app = if Array.is_empty lb_args
then lb_type_of_p else mkApp (lb_type_of_p,lb_args)
in
Tacticals.New.tclTHENLIST [
@@ -399,11 +405,12 @@ let do_replace_lb mode lb_scheme_key aavoid narg p q =
(* used in the bool -> leib side *)
let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
+ let open EConstr in
let avoid = Array.of_list aavoid in
- let do_arg v offset =
+ let do_arg sigma v offset =
try
let x = narg*offset in
- let s = destVar v in
+ let s = destVar sigma v in
let n = Array.length avoid in
let rec find i =
if Id.equal avoid.(n-i) s then avoid.(n-i-x)
@@ -416,7 +423,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
(* if this happen then the args have to be already declared as a
Parameter*)
(
- let mp,dir,lbl = repr_con (fst (destConst v)) in
+ let mp,dir,lbl = repr_con (fst (destConst sigma v)) in
mkConst (make_con mp dir (mk_label (
if Int.equal offset 1 then ("eq_"^(Label.to_string lbl))
else ((Label.to_string lbl)^"_bl")
@@ -429,9 +436,10 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
| (t1::q1,t2::q2) ->
Proofview.Goal.enter { enter = begin fun gl ->
let tt1 = Tacmach.New.pf_unsafe_type_of gl t1 in
- if eq_constr t1 t2 then aux q1 q2
+ let sigma = Tacmach.New.project gl in
+ if EConstr.eq_constr sigma t1 t2 then aux q1 q2
else (
- let u,v = try destruct_ind tt1
+ let u,v = try destruct_ind sigma tt1
(* trick so that the good sequence is returned*)
with e when CErrors.noncritical e -> indu,[||]
in if eq_ind (fst u) ind
@@ -439,7 +447,7 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
else (
let bl_t1, eff =
try
- let c, eff = find_scheme bl_scheme_key (out_punivs u) (*FIXME*) in
+ let c, eff = find_scheme bl_scheme_key (fst u) (*FIXME*) in
mkConst c, eff
with Not_found ->
(* spiwack: the format of this error message should probably
@@ -448,17 +456,17 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
(str "boolean->Leibniz:" ++
str "You have to declare the" ++
str "decidability over " ++
- Printer.pr_constr tt1 ++
+ Printer.pr_econstr tt1 ++
str " first.")
in
user_err err_msg
in let bl_args =
Array.append (Array.append
(Array.map (fun x -> x) v)
- (Array.map (fun x -> do_arg x 1) v))
- (Array.map (fun x -> do_arg x 2) v )
+ (Array.map (fun x -> do_arg sigma x 1) v))
+ (Array.map (fun x -> do_arg sigma x 2) v )
in
- let app = if Array.equal eq_constr bl_args [||]
+ let app = if Array.is_empty bl_args
then bl_t1 else mkApp (bl_t1,bl_args)
in
Tacticals.New.tclTHENLIST [
@@ -472,21 +480,22 @@ let do_replace_bl mode bl_scheme_key (ind,u as indu) aavoid narg lft rgt =
| ([],[]) -> Proofview.tclUNIT ()
| _ -> Tacticals.New.tclZEROMSG (str "Both side of the equality must have the same arity.")
in
- begin try Proofview.tclUNIT (destApp lft)
+ Proofview.tclEVARMAP >>= fun sigma ->
+ begin try Proofview.tclUNIT (destApp sigma lft)
with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.")
end >>= fun (ind1,ca1) ->
- begin try Proofview.tclUNIT (destApp rgt)
+ begin try Proofview.tclUNIT (destApp sigma rgt)
with DestKO -> Tacticals.New.tclZEROMSG (str "replace failed.")
end >>= fun (ind2,ca2) ->
- begin try Proofview.tclUNIT (out_punivs (destInd ind1))
+ begin try Proofview.tclUNIT (fst (destInd sigma ind1))
with DestKO ->
- begin try Proofview.tclUNIT (fst (fst (destConstruct ind1)))
+ begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind1)))
with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.")
end
end >>= fun (sp1,i1) ->
- begin try Proofview.tclUNIT (out_punivs (destInd ind2))
+ begin try Proofview.tclUNIT (fst (destInd sigma ind2))
with DestKO ->
- begin try Proofview.tclUNIT (fst (fst (destConstruct ind2)))
+ begin try Proofview.tclUNIT (fst (fst (destConstruct sigma ind2)))
with DestKO -> Tacticals.New.tclZEROMSG (str "The expected type is an inductive one.")
end
end >>= fun (sp2,i2) ->
@@ -518,7 +527,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 " ++ pr_mind (fst ind) ++ str " is needed.");
- in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)), eff
+ in (if Array.equal Term.eq_constr eA [||] then e else mkApp(e,eA)), eff
(**********************************************************************)
(* Boolean->Leibniz *)
@@ -573,12 +582,10 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec =
( List.map (fun (_,_,sbl,_ ) -> sbl) list_id )
in
let fresh_id s gl =
- Tacmach.New.of_old begin fun gsig ->
- let fresh = fresh_id (!avoid) s gsig in
+ let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in
avoid := fresh::(!avoid); fresh
- end gl
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -586,9 +593,9 @@ let compute_bl_tact mode bl_scheme_key ind lnamesparrec nparrec =
(* try with *)
Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
intro_using freshn ;
- induct_on (mkVar freshn);
+ induct_on (EConstr.mkVar freshn);
intro_using freshm;
- destruct_on (mkVar freshm);
+ destruct_on (EConstr.mkVar freshm);
intro_using freshz;
intros;
Tacticals.New.tclTRY (
@@ -600,10 +607,10 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
*)
Tacticals.New.tclREPEAT (
Tacticals.New.tclTHENLIST [
- Simple.apply_in freshz (andb_prop());
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Simple.apply_in freshz (EConstr.of_constr (andb_prop()));
+ Proofview.Goal.enter { enter = begin fun gl ->
let fresht = fresh_id (Id.of_string "Z") gl in
- destruct_on_as (mkVar freshz)
+ destruct_on_as (EConstr.mkVar freshz)
(IntroOrPattern [[dl,IntroNaming (IntroIdentifier fresht);
dl,IntroNaming (IntroIdentifier freshz)]])
end }
@@ -612,11 +619,12 @@ repeat ( apply andb_prop in z;let z1:= fresh "Z" in destruct z as [z1 z]).
Ci a1 ... an = Ci b1 ... bn
replace bi with ai; auto || replace bi with ai by apply typeofbi_prod ; auto
*)
- Proofview.Goal.nf_enter { enter = begin fun gls ->
- let gl = Proofview.Goal.concl gls in
- match (kind_of_term gl) with
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let concl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
+ match EConstr.kind sigma concl with
| App (c,ca) -> (
- match (kind_of_term c) with
+ match EConstr.kind sigma c with
| Ind (indeq, u) ->
if eq_gr (IndRef indeq) Coqlib.glob_eq
then
@@ -656,8 +664,9 @@ let make_bl_scheme mode mind =
let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in
let ctx = Evd.make_evar_universe_context (Global.env ()) None in
let side_eff = side_effect_of_mode mode in
+ let bl_goal = EConstr.of_constr bl_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx bl_goal
- (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, Univ.Instance.empty) lnamesparrec nparrec)
+ (compute_bl_tact mode (!bl_scheme_kind_aux()) (ind, EConstr.EInstance.empty) lnamesparrec nparrec)
in
([|ans|], ctx), eff
@@ -709,6 +718,7 @@ let compute_lb_goal ind lnamesparrec nparrec =
))), eff
let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
+ let open EConstr in
let list_id = list_id lnamesparrec in
let avoid = ref [] in
let first_intros =
@@ -717,12 +727,10 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
( List.map (fun (_,_,_,slb) -> slb) list_id )
in
let fresh_id s gl =
- Tacmach.New.of_old begin fun gsig ->
- let fresh = fresh_id (!avoid) s gsig in
+ let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in
avoid := fresh::(!avoid); fresh
- end gl
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -730,26 +738,27 @@ let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec =
(* try with *)
Tacticals.New.tclTHENLIST [ intros_using fresh_first_intros;
intro_using freshn ;
- induct_on (mkVar freshn);
+ induct_on (EConstr.mkVar freshn);
intro_using freshm;
- destruct_on (mkVar freshm);
+ destruct_on (EConstr.mkVar freshm);
intro_using freshz;
intros;
Tacticals.New.tclTRY (
Tacticals.New.tclORELSE reflexivity (Equality.discr_tac false None)
);
- Equality.inj None false None (mkVar freshz,NoBindings);
+ Equality.inj None false None (EConstr.mkVar freshz,NoBindings);
intros; simpl_in_concl;
Auto.default_auto;
Tacticals.New.tclREPEAT (
- Tacticals.New.tclTHENLIST [apply (andb_true_intro());
+ Tacticals.New.tclTHENLIST [apply (EConstr.of_constr (andb_true_intro()));
simplest_split ;Auto.default_auto ]
);
- Proofview.Goal.nf_enter { enter = begin fun gls ->
- let gl = Proofview.Goal.concl gls in
+ Proofview.Goal.enter { enter = begin fun gls ->
+ let concl = Proofview.Goal.concl gls in
+ let sigma = Tacmach.New.project gl in
(* assume the goal to be eq (eq_type ...) = true *)
- match (kind_of_term gl) with
- | App(c,ca) -> (match (kind_of_term ca.(1)) with
+ match EConstr.kind sigma concl with
+ | App(c,ca) -> (match (EConstr.kind sigma ca.(1)) with
| App(c',ca') ->
let n = Array.length ca' in
do_replace_lb mode lb_scheme_key
@@ -780,6 +789,7 @@ let make_lb_scheme mode mind =
let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in
let ctx = Evd.make_evar_universe_context (Global.env ()) None in
let side_eff = side_effect_of_mode mode in
+ let lb_goal = EConstr.of_constr lb_goal in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx lb_goal
(compute_lb_tact mode (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)
in
@@ -865,12 +875,10 @@ let compute_dec_tact ind lnamesparrec nparrec =
( List.map (fun (_,_,_,slb) -> slb) list_id )
in
let fresh_id s gl =
- Tacmach.New.of_old begin fun gsig ->
- let fresh = fresh_id (!avoid) s gsig in
+ let fresh = fresh_id_in_env (!avoid) s (Proofview.Goal.env gl) in
avoid := fresh::(!avoid); fresh
- end gl
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let fresh_first_intros = List.map (fun id -> fresh_id id gl) first_intros in
let freshn = fresh_id (Id.of_string "x") gl in
let freshm = fresh_id (Id.of_string "y") gl in
@@ -896,24 +904,24 @@ let compute_dec_tact ind lnamesparrec nparrec =
intros_using fresh_first_intros;
intros_using [freshn;freshm];
(*we do this so we don't have to prove the same goal twice *)
- assert_by (Name freshH) (
+ assert_by (Name freshH) (EConstr.of_constr (
mkApp(sumbool(),[|eqtrue eqbnm; eqfalse eqbnm|])
- )
- (Tacticals.New.tclTHEN (destruct_on eqbnm) Auto.default_auto);
+ ))
+ (Tacticals.New.tclTHEN (destruct_on (EConstr.of_constr eqbnm)) Auto.default_auto);
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let freshH2 = fresh_id (Id.of_string "H") gl in
- Tacticals.New.tclTHENS (destruct_on_using (mkVar freshH) freshH2) [
+ Tacticals.New.tclTHENS (destruct_on_using (EConstr.mkVar freshH) freshH2) [
(* left *)
Tacticals.New.tclTHENLIST [
simplest_left;
- apply (mkApp(blI,Array.map(fun x->mkVar x) xargs));
+ apply (EConstr.of_constr (mkApp(blI,Array.map(fun x->mkVar x) xargs)));
Auto.default_auto
]
;
(*right *)
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let freshH3 = fresh_id (Id.of_string "H") gl in
Tacticals.New.tclTHENLIST [
simplest_right ;
@@ -921,15 +929,15 @@ let compute_dec_tact ind lnamesparrec nparrec =
intro;
Equality.subst_all ();
assert_by (Name freshH3)
- (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|]))
+ (EConstr.of_constr (mkApp(eq,[|bb;mkApp(eqI,[|mkVar freshm;mkVar freshm|]);tt|])))
(Tacticals.New.tclTHENLIST [
- apply (mkApp(lbI,Array.map (fun x->mkVar x) xargs));
+ apply (EConstr.of_constr (mkApp(lbI,Array.map (fun x->mkVar x) xargs)));
Auto.default_auto
]);
Equality.general_rewrite_bindings_in true
Locus.AllOccurrences true false
(List.hd !avoid)
- ((mkVar (List.hd (List.tl !avoid))),
+ ((EConstr.mkVar (List.hd (List.tl !avoid))),
NoBindings
)
true;
@@ -954,7 +962,7 @@ let make_eq_decidability mode mind =
context_chop (nparams-nparrec) mib.mind_params_ctxt in
let side_eff = side_effect_of_mode mode in
let (ans, _, ctx) = Pfedit.build_by_tactic ~side_eff (Global.env()) ctx
- (compute_dec_goal (ind,u) lnamesparrec nparrec)
+ (EConstr.of_constr (compute_dec_goal (ind,u) lnamesparrec nparrec))
(compute_dec_tact ind lnamesparrec nparrec)
in
([|ans|], ctx), Safe_typing.empty_private_constants
diff --git a/vernac/class.ml b/vernac/class.ml
index 0dc799014..95114ec39 100644
--- a/vernac/class.ml
+++ b/vernac/class.ml
@@ -62,7 +62,7 @@ let explain_coercion_error g = function
(* Verifications pour l'ajout d'une classe *)
let check_reference_arity ref =
- if not (Reductionops.is_arity (Global.env()) Evd.empty (Global.type_of_global_unsafe ref)) then
+ if not (Reductionops.is_arity (Global.env()) Evd.empty (EConstr.of_constr (Global.type_of_global_unsafe ref))) (** FIXME *) then
raise (CoercionError (NotAClass ref))
let check_arity = function
@@ -81,12 +81,13 @@ let check_target clt = function
(* condition d'heritage uniforme *)
-let uniform_cond nargs lt =
+let uniform_cond sigma nargs lt =
+ let open EConstr in
let rec aux = function
| (0,[]) -> true
| (n,t::l) ->
- let t = strip_outer_cast t in
- isRel t && Int.equal (destRel t) n && aux ((n-1),l)
+ let t = strip_outer_cast sigma t in
+ isRel sigma t && Int.equal (destRel sigma t) n && aux ((n-1),l)
| _ -> false
in
aux (nargs,lt)
@@ -120,7 +121,7 @@ let get_source lp source =
let (cl1,u1,lv1) =
match lp with
| [] -> raise Not_found
- | t1::_ -> find_class_type Evd.empty t1
+ | t1::_ -> find_class_type Evd.empty (EConstr.of_constr t1)
in
(cl1,u1,lv1,1)
| Some cl ->
@@ -128,7 +129,7 @@ let get_source lp source =
| [] -> raise Not_found
| t1::lt ->
try
- let cl1,u1,lv1 = find_class_type Evd.empty t1 in
+ let cl1,u1,lv1 = find_class_type Evd.empty (EConstr.of_constr t1) in
if cl_typ_eq cl cl1 then cl1,u1,lv1,(List.length lt+1)
else raise Not_found
with Not_found -> aux lt
@@ -138,7 +139,7 @@ let get_target t ind =
if (ind > 1) then
CL_FUN
else
- match pi1 (find_class_type Evd.empty t) with
+ match pi1 (find_class_type Evd.empty (EConstr.of_constr t)) with
| CL_CONST p when Environ.is_projection p (Global.env ()) ->
CL_PROJ p
| x -> x
@@ -193,20 +194,20 @@ let build_id_coercion idf_opt source poly =
let val_f =
it_mkLambda_or_LetIn
(mkLambda (Name Namegen.default_dependent_ident,
- applistc vs (Context.Rel.to_extended_list 0 lams),
+ applistc vs (Context.Rel.to_extended_list mkRel 0 lams),
mkRel 1))
lams
in
let typ_f =
- it_mkProd_wo_LetIn
- (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list 0 lams), lift 1 t))
+ List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d)
+ (mkProd (Anonymous, applistc vs (Context.Rel.to_extended_list mkRel 0 lams), lift 1 t))
lams
in
(* juste pour verification *)
let _ =
if not
(Reductionops.is_conv_leq env sigma
- (Typing.unsafe_type_of env sigma val_f) typ_f)
+ (Typing.unsafe_type_of env sigma (EConstr.of_constr val_f)) (EConstr.of_constr typ_f))
then
user_err (strbrk
"Cannot be defined as coercion (maybe a bad number of arguments).")
@@ -215,7 +216,7 @@ let build_id_coercion idf_opt source poly =
match idf_opt with
| Some idf -> idf
| None ->
- let cl,u,_ = find_class_type sigma t in
+ let cl,u,_ = find_class_type sigma (EConstr.of_constr t) in
Id.of_string ("Id_"^(ident_key_of_class source)^"_"^
(ident_key_of_class cl))
in
@@ -263,7 +264,7 @@ let add_new_coercion_core coef stre poly source target isid =
raise (CoercionError (NoSource source))
in
check_source (Some cls);
- if not (uniform_cond (llp-ind) lvs) then
+ if not (uniform_cond Evd.empty (** FIXME *) (llp-ind) lvs) then
warn_uniform_inheritance coef;
let clt =
try
diff --git a/vernac/classes.ml b/vernac/classes.ml
index c577fe6e3..833719965 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -49,7 +49,7 @@ let set_typeclass_transparency c local b =
let _ =
Hook.set Typeclasses.add_instance_hint_hook
(fun inst path local info poly ->
- let inst' = match inst with IsConstr c -> Hints.IsConstr (c, Univ.ContextSet.empty)
+ let inst' = match inst with IsConstr c -> Hints.IsConstr (EConstr.of_constr c, Univ.ContextSet.empty)
| IsGlobal gr -> Hints.IsGlobRef gr
in
let info =
@@ -74,7 +74,7 @@ let existing_instance glob g info =
let info = Option.default Hints.empty_hint_info info in
let instance = Global.type_of_global_unsafe c in
let _, r = decompose_prod_assum instance in
- match class_of_constr r with
+ match class_of_constr Evd.empty (EConstr.of_constr r) with
| Some (_, ((tc,u), _)) -> add_instance (new_instance tc info glob
(*FIXME*) (Flags.use_polymorphic_flag ()) c)
| None -> user_err ~loc:(loc_of_reference g)
@@ -92,7 +92,7 @@ let type_ctx_instance evars env ctx inst subst =
let t' = substl subst (RelDecl.get_type decl) in
let c', l =
match decl with
- | LocalAssum _ -> interp_casted_constr_evars env evars (List.hd l) t', List.tl l
+ | LocalAssum _ -> EConstr.Unsafe.to_constr (interp_casted_constr_evars env evars (List.hd l) t'), List.tl l
| LocalDef (_,b,_) -> substl subst b, l
in
let d = RelDecl.get_name decl, Some c', t' in
@@ -159,13 +159,16 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
in
let k, u, cty, ctx', ctx, len, imps, subst =
let impls, ((env', ctx), imps) = interp_context_evars env evars ctx in
+ let ctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) ctx in
let c', imps' = interp_type_evars_impls ~impls env' evars tclass in
+ let c' = EConstr.Unsafe.to_constr c' in
let len = List.length ctx in
let imps = imps @ Impargs.lift_implicits len imps' in
let ctx', c = decompose_prod_assum c' in
let ctx'' = ctx' @ ctx in
- let k, args = Typeclasses.dest_class_app (push_rel_context ctx'' env) c in
- let cl, u = Typeclasses.typeclass_univ_instance k in
+ let (k, u), args = Typeclasses.dest_class_app (push_rel_context ctx'' env) !evars (EConstr.of_constr c) in
+ let u = EConstr.EInstance.kind !evars u in
+ let cl, u = Typeclasses.typeclass_univ_instance (k, u) in
let _, args =
List.fold_right (fun decl (args, args') ->
match decl with
@@ -189,7 +192,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
let env' = push_rel_context ctx env in
evars := Evarutil.nf_evar_map !evars;
evars := resolve_typeclasses ~filter:Typeclasses.all_evars ~fail:true env !evars;
- let subst = List.map (Evarutil.nf_evar !evars) subst in
+ let subst = List.map (fun c -> EConstr.Unsafe.to_constr (Evarutil.nf_evar !evars (EConstr.of_constr c))) subst in
if abstract then
begin
let subst = List.fold_left2
@@ -202,7 +205,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
let t = it_mkProd_or_LetIn ty_constr (ctx' @ ctx) in
nf t
in
- Pretyping.check_evars env Evd.empty !evars termtype;
+ Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr termtype);
let pl, ctx = Evd.universe_context ?names:pl !evars in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id
(ParameterEntry
@@ -228,6 +231,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
| None -> if List.is_empty k.cl_props then Some (Inl subst) else None
| Some (Inr term) ->
let c = interp_casted_constr_evars env' evars term cty in
+ let c = EConstr.Unsafe.to_constr c in
Some (Inr (c, subst))
| Some (Inl props) ->
let get_id =
@@ -298,7 +302,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
let evm, nf = Evarutil.nf_evar_map_universes !evars in
let termtype = nf termtype in
let _ = (* Check that the type is free of evars now. *)
- Pretyping.check_evars env Evd.empty evm termtype
+ Pretyping.check_evars env Evd.empty evm (EConstr.of_constr termtype)
in
let term = Option.map nf term in
if not (Evd.has_undefined evm) && not (Option.is_empty term) then
@@ -334,14 +338,14 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
the refinement manually.*)
let gls = List.rev (Evd.future_goals evm) in
let evm = Evd.reset_future_goals evm in
- Lemmas.start_proof id ?pl kind evm termtype
+ Lemmas.start_proof id ?pl kind evm (EConstr.of_constr termtype)
(Lemmas.mk_hook
(fun _ -> instance_hook k pri global imps ?hook));
(* spiwack: I don't know what to do with the status here. *)
if not (Option.is_empty term) then
let init_refine =
Tacticals.New.tclTHENLIST [
- Refine.refine { run = fun evm -> Sigma (Option.get term, evm, Sigma.refl) };
+ Refine.refine { run = fun evm -> Sigma (EConstr.of_constr (Option.get term), evm, Sigma.refl) };
Proofview.Unsafe.tclNEWGOALS gls;
Tactics.New.reduce_after_refine;
]
@@ -370,9 +374,10 @@ let context poly l =
let env = Global.env() in
let evars = ref (Evd.from_env env) in
let _, ((env', fullctx), impls) = interp_context_evars env evars l in
+ let fullctx = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) fullctx in
let subst = Evarutil.evd_comb0 Evarutil.nf_evars_and_universes evars in
let fullctx = Context.Rel.map subst fullctx in
- let ce t = Pretyping.check_evars env Evd.empty !evars t in
+ let ce t = Pretyping.check_evars env Evd.empty !evars (EConstr.of_constr t) in
let () = List.iter (fun decl -> Context.Rel.Declaration.iter_constr ce decl) fullctx in
let ctx =
try named_of_rel_context fullctx
@@ -387,7 +392,7 @@ let context poly l =
let () = uctx := Univ.ContextSet.empty in
let decl = (ParameterEntry (None,poly,(t,ctx),None), IsAssumption Logical) in
let cst = Declare.declare_constant ~internal:Declare.InternalTacticRequest id decl in
- match class_of_constr t with
+ match class_of_constr !evars (EConstr.of_constr t) with
| Some (rels, ((tc,_), args) as _cl) ->
add_instance (Typeclasses.new_instance tc Hints.empty_hint_info false (*FIXME*)
poly (ConstRef cst));
diff --git a/vernac/command.ml b/vernac/command.ml
index 6eb7037f8..781bf3223 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -45,7 +45,7 @@ let do_universe poly l = Declare.do_universe poly l
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 c else
+ if Int.equal n 0 then f env sigma (EConstr.of_constr c) else
match kind_of_term c with
| Lambda (x,t,c) ->
mkLambda (x,t,under_binders (push_rel (LocalAssum (x,t)) env) sigma f (n-1) c)
@@ -79,7 +79,7 @@ let red_constant_entry n ce sigma = function
let redfun env sigma c =
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (c, _, _) = redfun.e_redfun env sigma c in
- c
+ EConstr.Unsafe.to_constr c
in
{ ce with const_entry_body = Future.chain ~pure:true proof_out
(fun ((body,ctx),eff) -> (under_binders env sigma redfun n body,ctx),eff) }
@@ -95,6 +95,7 @@ let interp_definition pl bl p red_option c ctypopt =
let ctx = Evd.make_evar_universe_context env pl in
let evdref = ref (Evd.from_ctx ctx) in
let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in
+ let ctx = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx in
let nb_args = List.length ctx in
let imps,pl,ce =
match ctypopt with
@@ -103,6 +104,7 @@ let interp_definition pl bl p red_option c ctypopt =
let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
let env_bl = push_rel_context ctx env in
let c, imps2 = interp_constr_evars_impls ~impls env_bl evdref c in
+ let c = EConstr.Unsafe.to_constr c in
let nf,subst = Evarutil.e_nf_evars_and_universes evdref in
let body = nf (it_mkLambda_or_LetIn c ctx) in
let vars = Universes.universes_of_constr body in
@@ -116,9 +118,11 @@ let interp_definition pl bl p red_option c ctypopt =
let ctx = Context.Rel.map (Vars.subst_univs_constr subst) ctx in
let env_bl = push_rel_context ctx env in
let c, imps2 = interp_casted_constr_evars_impls ~impls env_bl evdref c ty in
+ let c = EConstr.Unsafe.to_constr c in
let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
let body = nf (it_mkLambda_or_LetIn c ctx) in
- let typ = nf (it_mkProd_or_LetIn ty ctx) in
+ let ty = EConstr.Unsafe.to_constr ty in
+ let typ = nf (Term.it_mkProd_or_LetIn ty ctx) in
let beq b1 b2 = if b1 then b2 else not b2 in
let impl_eq (x,y,z) (x',y',z') = beq x x' && beq y y' && beq z z' in
(* Check that all implicit arguments inferable from the term
@@ -206,7 +210,7 @@ let do_definition ident k pl bl red_option c ctypopt hook =
assert(Univ.ContextSet.is_empty ctx);
let typ = match ce.const_entry_type with
| Some t -> t
- | None -> Retyping.get_type_of env evd c
+ | None -> EConstr.Unsafe.to_constr (Retyping.get_type_of env evd (EConstr.of_constr c))
in
Obligations.check_evars env evd;
let obls, _, c, cty =
@@ -263,7 +267,9 @@ match local with
let interp_assumption evdref env impls bl c =
let c = prod_constr_expr c bl in
- interp_type_evars_impls env evdref ~impls c
+ let ty, impls = interp_type_evars_impls env evdref ~impls c in
+ let ty = EConstr.Unsafe.to_constr ty in
+ (ty, impls)
let declare_assumptions idl is_coe k (c,ctx) pl imps impl_is_on nl =
let refs, status, _ =
@@ -303,7 +309,8 @@ let do_assumptions_unbound_univs (_, poly, _ as kind) nl l =
(* The universe constraints come from the whole telescope. *)
let evd = Evd.nf_constraints evd in
let ctx = Evd.universe_context_set evd in
- let l = List.map (on_pi2 (nf_evar evd)) l in
+ let nf_evar c = EConstr.Unsafe.to_constr (nf_evar evd (EConstr.of_constr c)) in
+ let l = List.map (on_pi2 nf_evar) l in
pi2 (List.fold_left (fun (subst,status,ctx) ((is_coe,idl),t,imps) ->
let t = replace_vars subst t in
let (refs,status') = declare_assumptions idl is_coe kind (t,ctx) [] imps false nl in
@@ -321,6 +328,7 @@ let do_assumptions_bound_univs coe kind nl id pl c =
let evdref = ref (Evd.from_ctx ctx) in
let ty, impls = interp_type_evars_impls env evdref c in
let nf, subst = Evarutil.e_nf_evars_and_universes evdref in
+ let ty = EConstr.Unsafe.to_constr ty in
let ty = nf ty in
let vars = Universes.universes_of_constr ty in
let evd = Evd.restrict_universe_context !evdref vars in
@@ -398,7 +406,7 @@ let check_all_names_different indl =
| _ -> raise (InductiveError (SameNamesOverlap l))
let mk_mltype_data evdref env assums arity indname =
- let is_ml_type = is_sort env !evdref arity in
+ let is_ml_type = is_sort env !evdref (EConstr.of_constr arity) in
(is_ml_type,indname,assums)
let prepare_param = function
@@ -442,9 +450,10 @@ let interp_ind_arity env evdref ind =
let imps = Implicit_quantifiers.implicits_of_glob_constr ~with_products:true c in
let t, impls = understand_tcc_evars env evdref ~expected_type:IsType c, imps in
let pseudo_poly = check_anonymous_type c in
- let () = if not (Reduction.is_arity env t) then
+ let () = if not (Reductionops.is_arity env !evdref t) then
user_err ~loc:(constr_loc ind.ind_arity) (str "Not an arity")
in
+ let t = EConstr.Unsafe.to_constr t in
t, pseudo_poly, impls
let interp_cstrs evdref env impls mldata arity ind =
@@ -452,7 +461,7 @@ let interp_cstrs evdref env impls mldata arity ind =
(* Complete conclusions of constructor types if given in ML-style syntax *)
let ctyps' = List.map2 (complete_conclusion mldata) cnames ctyps in
(* Interpret the constructor types *)
- let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls) ctyps') in
+ let ctyps'', cimpls = List.split (List.map (interp_type_evars_impls evdref env ~impls %> on_fst EConstr.Unsafe.to_constr) ctyps') in
(cnames, ctyps'', cimpls)
let sign_level env evd sign =
@@ -462,7 +471,7 @@ let sign_level env evd sign =
| LocalDef _ -> lev, push_rel d env
| LocalAssum _ ->
let s = destSort (Reduction.whd_all env
- (nf_evar evd (Retyping.get_type_of env evd (RelDecl.get_type d))))
+ (EConstr.Unsafe.to_constr (nf_evar evd (Retyping.get_type_of env evd (EConstr.of_constr (RelDecl.get_type d))))))
in
let u = univ_of_sort s in
(Univ.sup u lev, push_rel d env))
@@ -575,6 +584,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
let _, ((env_params, ctx_params), userimpls) =
interp_context_evars env0 evdref paramsl
in
+ let ctx_params = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx_params in
let indnames = List.map (fun ind -> ind.ind_name) indl in
(* Names of parameters as arguments of the inductive type (defs removed) *)
@@ -584,7 +594,7 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
(* Interpret the arities *)
let arities = List.map (interp_ind_arity env_params evdref) indl in
- let fullarities = List.map (fun (c, _, _) -> it_mkProd_or_LetIn c ctx_params) arities in
+ let fullarities = List.map (fun (c, _, _) -> Term.it_mkProd_or_LetIn c ctx_params) arities in
let env_ar = push_types env0 indnames fullarities in
let env_ar_params = push_rel_context ctx_params env_ar in
@@ -619,10 +629,10 @@ let interp_mutual_inductive (paramsl,indl) notations poly prv finite =
let ctx_params = Context.Rel.map nf ctx_params in
let evd = !evdref in
let pl, uctx = Evd.universe_context ?names:pl evd in
- List.iter (check_evars env_params Evd.empty evd) arities;
- Context.Rel.iter (check_evars env0 Evd.empty evd) ctx_params;
+ List.iter (fun c -> check_evars env_params Evd.empty evd (EConstr.of_constr c)) arities;
+ Context.Rel.iter (fun c -> check_evars env0 Evd.empty evd (EConstr.of_constr c)) ctx_params;
List.iter (fun (_,ctyps,_) ->
- List.iter (check_evars env_ar_params Evd.empty evd) ctyps)
+ List.iter (fun c -> check_evars env_ar_params Evd.empty evd (EConstr.of_constr c)) ctyps)
constructors;
(* Build the inductive entries *)
@@ -686,7 +696,7 @@ let is_recursive mie =
let rec is_recursive_constructor lift typ =
match Term.kind_of_term typ with
| Prod (_,arg,rest) ->
- Termops.dependent (mkRel lift) arg ||
+ not (EConstr.Vars.noccurn Evd.empty (** FIXME *) lift (EConstr.of_constr arg)) ||
is_recursive_constructor (lift+1) rest
| LetIn (na,b,t,rest) -> is_recursive_constructor (lift+1) rest
| _ -> false
@@ -814,11 +824,11 @@ let warn_non_full_mutual =
(fun (x,xge,y,yge,isfix,rest) ->
non_full_mutual_message x xge y yge isfix rest)
-let check_mutuality env isfix fixl =
+let check_mutuality env evd isfix fixl =
let names = List.map fst fixl in
let preorder =
List.map (fun (id,def) ->
- (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env id' def) names))
+ (id, List.filter (fun id' -> not (Id.equal id id') && occur_var env evd id' (EConstr.of_constr def)) names))
fixl in
let po = partial_order Id.equal preorder in
match List.filter (function (_,Inr _) -> true | _ -> false) po with
@@ -843,15 +853,17 @@ let interp_fix_context env evdref isfix fix =
((env'', ctx' @ ctx), (impl_env',imps @ imps'), annot)
let interp_fix_ccl evdref impls (env,_) fix =
- interp_type_evars_impls ~impls env evdref fix.fix_type
+ let (c, impl) = interp_type_evars_impls ~impls env evdref fix.fix_type in
+ (c, impl)
let interp_fix_body env_rec evdref impls (_,ctx) fix ccl =
+ let open EConstr in
Option.map (fun body ->
let env = push_rel_context ctx env_rec in
let body = interp_casted_constr_evars env evdref ~impls body ccl in
it_mkLambda_or_LetIn body ctx) fix.fix_body
-let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx
+let build_fix_type (_,ctx) ccl = EConstr.it_mkProd_or_LetIn ccl ctx
let declare_fix ?(opaque = false) (_,poly,_ as kind) pl ctx f ((def,_),eff) t imps =
let ce = definition_entry ~opaque ~types:t ~poly ~univs:ctx ~eff def in
@@ -891,21 +903,25 @@ let fixsub_module = subtac_dir @ ["Wf"]
let tactics_module = subtac_dir @ ["Tactics"]
let init_reference dir s () = Coqlib.gen_reference "Command" dir s
-let init_constant dir s () = Coqlib.gen_constant "Command" dir s
+let init_constant dir s () = EConstr.of_constr (Coqlib.gen_constant "Command" dir s)
let make_ref l s = init_reference l s
let fix_proto = init_constant tactics_module "fix_proto"
let fix_sub_ref = make_ref fixsub_module "Fix_sub"
let measure_on_R_ref = make_ref fixsub_module "MR"
let well_founded = init_constant ["Init"; "Wf"] "well_founded"
let mkSubset name typ prop =
- mkApp (Universes.constr_of_global (delayed_force build_sigma).typ,
+ let open EConstr in
+ mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).typ),
[| typ; mkLambda (name, typ, prop) |])
let sigT = Lazy.from_fun build_sigma_type
let make_qref s = Qualid (Loc.ghost, qualid_of_string s)
let lt_ref = make_qref "Init.Peano.lt"
-let rec telescope = function
+let rec telescope l =
+ let open EConstr in
+ let open Vars in
+ match l with
| [] -> assert false
| [LocalAssum (n, t)] -> t, [LocalDef (n, mkRel 1, t)], mkRel 1
| LocalAssum (n, t) :: tl ->
@@ -915,7 +931,9 @@ let rec telescope = function
let t = RelDecl.get_type decl in
let pred = mkLambda (RelDecl.get_name decl, t, ty) in
let ty = Universes.constr_of_global (Lazy.force sigT).typ in
+ let ty = EConstr.of_constr ty in
let intro = Universes.constr_of_global (Lazy.force sigT).intro in
+ let intro = EConstr.of_constr intro in
let sigty = mkApp (ty, [|t; pred|]) in
let intro = mkApp (intro, [|lift k t; lift k pred; mkRel k; constr|]) in
(sigty, pred :: tys, (succ k, intro)))
@@ -925,9 +943,11 @@ let rec telescope = function
(fun pred decl (prev, subst) ->
let t = RelDecl.get_type decl in
let p1 = Universes.constr_of_global (Lazy.force sigT).proj1 in
+ let p1 = EConstr.of_constr p1 in
let p2 = Universes.constr_of_global (Lazy.force sigT).proj2 in
- let proj1 = applistc p1 [t; pred; prev] in
- let proj2 = applistc p2 [t; pred; prev] in
+ let p2 = EConstr.of_constr p2 in
+ let proj1 = applist (p1, [t; pred; prev]) in
+ let proj2 = applist (p2, [t; pred; prev]) in
(lift 1 proj2, LocalDef (get_name decl, proj1, t) :: subst))
(List.rev tys) tl (mkRel 1, [])
in ty, (LocalDef (n, last, t) :: subst), constr
@@ -936,9 +956,12 @@ let rec telescope = function
ty, (LocalDef (n, b, t) :: subst), lift 1 term
let nf_evar_context sigma ctx =
- List.map (map_constr (Evarutil.nf_evar sigma)) ctx
+ List.map (map_constr (fun c -> Evarutil.nf_evar sigma c)) ctx
let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
+ let open EConstr in
+ let open Vars in
+ let lift_rel_context n l = Termops.map_rel_context_with_binders (liftn n) l in
Coqlib.check_required_library ["Coq";"Program";"Wf"];
let env = Global.env() in
let ctx = Evd.make_evar_universe_context env pl in
@@ -959,23 +982,25 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let error () =
user_err ~loc:(constr_loc r)
~hdr:"Command.build_wellfounded"
- (Printer.pr_constr_env env !evdref rel ++ str " is not an homogeneous binary relation.")
+ (Printer.pr_econstr_env env !evdref rel ++ str " is not an homogeneous binary relation.")
in
try
let ctx, ar = Reductionops.splay_prod_n env !evdref 2 relty in
- match ctx, kind_of_term ar with
- | [LocalAssum (_,t); LocalAssum (_,u)], Sort (Prop Null)
- when Reductionops.is_conv env !evdref t u -> t
+ match ctx, EConstr.kind !evdref ar with
+ | [LocalAssum (_,t); LocalAssum (_,u)], Sort s
+ when Sorts.is_prop (ESorts.kind !evdref s) && Reductionops.is_conv env !evdref t u -> t
| _, _ -> error ()
with e when CErrors.noncritical e -> error ()
in
+ let relargty = EConstr.Unsafe.to_constr relargty in
let measure = interp_casted_constr_evars binders_env evdref measure relargty in
let wf_rel, wf_rel_fun, measure_fn =
let measure_body, measure =
it_mkLambda_or_LetIn measure letbinders,
it_mkLambda_or_LetIn measure binders
in
- let comb = Universes.constr_of_global (delayed_force measure_on_R_ref) in
+ let comb = EConstr.of_constr (Universes.constr_of_global (delayed_force measure_on_R_ref)) in
+ let relargty = EConstr.of_constr relargty in
let wf_rel = mkApp (comb, [| argtyp; relargty; rel; measure |]) in
let wf_rel_fun x y =
mkApp (rel, [| subst1 x measure_body;
@@ -990,7 +1015,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
in
let intern_bl = wfarg 1 :: [arg] in
let _intern_env = push_rel_context intern_bl env in
- let proj = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1 in
+ let proj = (*FIXME*)EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).Coqlib.proj1) in
let wfargpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel 3)) in
let projection = (* in wfarg :: arg :: before *)
mkApp (proj, [| argtyp ; wfargpred ; mkRel 1 |])
@@ -1003,7 +1028,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let intern_fun_binder = LocalAssum (Name (add_suffix recname "'"), intern_fun_arity_prod) in
let curry_fun =
let wfpred = mkLambda (Name argid', argtyp, wf_rel_fun (mkRel 1) (mkRel (2 * len + 4))) in
- let intro = (*FIXME*)Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro in
+ let intro = (*FIXME*)EConstr.of_constr (Universes.constr_of_global (delayed_force build_sigma).Coqlib.intro) in
let arg = mkApp (intro, [| argtyp; wfpred; lift 1 make; mkRel 1 |]) in
let app = mkApp (mkRel (2 * len + 2 (* recproof + orig binders + current binders *)), [| arg |]) in
let rcurry = mkApp (rel, [| measure; lift len measure |]) in
@@ -1013,23 +1038,23 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
LocalDef (Name recname, body, ty)
in
let fun_bl = intern_fun_binder :: [arg] in
- let lift_lets = Termops.lift_rel_context 1 letbinders in
+ let lift_lets = lift_rel_context 1 letbinders in
let intern_body =
let ctx = LocalAssum (Name recname, get_type curry_fun) :: binders_rel in
let (r, l, impls, scopes) =
Constrintern.compute_internalization_data env
- Constrintern.Recursive full_arity impls
+ Constrintern.Recursive (EConstr.Unsafe.to_constr full_arity) impls
in
let newimpls = Id.Map.singleton recname
(r, l, impls @ [(Some (Id.of_string "recproof", Impargs.Manual, (true, false)))],
scopes @ [None]) in
interp_casted_constr_evars (push_rel_context ctx env) evdref
- ~impls:newimpls body (lift 1 top_arity)
+ ~impls:newimpls body (EConstr.Unsafe.to_constr (lift 1 top_arity))
in
let intern_body_lam = it_mkLambda_or_LetIn intern_body (curry_fun :: lift_lets @ fun_bl) in
let prop = mkLambda (Name argname, argtyp, top_arity_let) in
let def =
- mkApp (Universes.constr_of_global (delayed_force fix_sub_ref),
+ mkApp (EConstr.of_constr (Universes.constr_of_global (delayed_force fix_sub_ref)),
[| argtyp ; wf_rel ;
Evarutil.e_new_evar env evdref
~src:(Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false)) wf_proof;
@@ -1045,11 +1070,12 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
if List.length binders_rel > 1 then
let name = add_suffix recname "_func" in
let hook l gr _ =
- let body = it_mkLambda_or_LetIn (mkApp (Universes.constr_of_global gr, [|make|])) binders_rel in
+ let body = it_mkLambda_or_LetIn (mkApp (EConstr.of_constr (Universes.constr_of_global gr), [|make|])) binders_rel in
let ty = it_mkProd_or_LetIn top_arity binders_rel in
+ let ty = EConstr.Unsafe.to_constr ty in
let pl, univs = Evd.universe_context ?names:pl !evdref in
(*FIXME poly? *)
- let ce = definition_entry ~poly ~types:ty ~univs (Evarutil.nf_evar !evdref body) in
+ let ce = definition_entry ~poly ~types:ty ~univs (EConstr.Unsafe.to_constr (Evarutil.nf_evar !evdref body)) in
(** FIXME: include locality *)
let c = Declare.declare_constant recname (DefinitionEntry ce, IsDefinition Definition) in
let gr = ConstRef c in
@@ -1068,6 +1094,8 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let hook = Lemmas.mk_hook hook in
let fullcoqc = Evarutil.nf_evar !evdref def in
let fullctyp = Evarutil.nf_evar !evdref typ in
+ let fullcoqc = EConstr.Unsafe.to_constr fullcoqc in
+ let fullctyp = EConstr.Unsafe.to_constr fullctyp in
Obligations.check_evars env !evdref;
let evars, _, evars_def, evars_typ =
Obligations.eterm_obligations env recname !evdref 0 fullcoqc fullctyp
@@ -1078,6 +1106,7 @@ let build_wellfounded (recname,pl,n,bl,arityc,body) poly r measure notation =
let interp_recursive isfix fixl notations =
let open Context.Named.Declaration in
+ let open EConstr in
let env = Global.env() in
let fixnames = List.map (fun fix -> fix.fix_name) fixl in
@@ -1098,14 +1127,14 @@ let interp_recursive isfix fixl notations =
let fixctximpenvs, fixctximps = List.split fiximppairs in
let fixccls,fixcclimps = List.split (List.map3 (interp_fix_ccl evdref) fixctximpenvs fixctxs fixl) in
let fixtypes = List.map2 build_fix_type fixctxs fixccls in
- let fixtypes = List.map (nf_evar !evdref) fixtypes in
+ let fixtypes = List.map (fun c -> nf_evar !evdref c) fixtypes in
let fiximps = List.map3
(fun ctximps cclimps (_,ctx) -> ctximps@(Impargs.lift_implicits (List.length ctx) cclimps))
fixctximps fixcclimps fixctxs in
let rec_sign =
List.fold_left2
(fun env' id t ->
- if Flags.is_program_mode () then
+ if Flags.is_program_mode () then
let sort = Evarutil.evd_comb1 (Typing.type_of ~refresh:true env) evdref t in
let fixprot =
try
@@ -1120,6 +1149,8 @@ let interp_recursive isfix fixl notations =
let env_rec = push_named_context rec_sign env in
(* Get interpretation metadatas *)
+ let fixtypes = List.map EConstr.Unsafe.to_constr fixtypes in
+ let fixccls = List.map EConstr.Unsafe.to_constr fixccls in
let impls = compute_internalization_env env Recursive fixnames fixtypes fiximps in
(* Interp bodies with rollback because temp use of notations/implicit *)
@@ -1134,6 +1165,7 @@ let interp_recursive isfix fixl notations =
(* Instantiate evars and check all are resolved *)
let evd = solve_unif_constraints_with_heuristics env_rec !evdref in
let evd, nf = nf_evars_and_universes evd in
+ let fixdefs = List.map (fun c -> Option.map EConstr.Unsafe.to_constr c) fixdefs in
let fixdefs = List.map (Option.map nf) fixdefs in
let fixtypes = List.map nf fixtypes in
let fixctxnames = List.map (fun (_,ctx) -> List.map RelDecl.get_name ctx) fixctxs in
@@ -1145,7 +1177,7 @@ let check_recursive isfix env evd (fixnames,fixdefs,_) =
check_evars_are_solved env evd Evd.empty;
if List.for_all Option.has_some fixdefs then begin
let fixdefs = List.map Option.get fixdefs in
- check_mutuality env isfix (List.combine fixnames fixdefs)
+ check_mutuality env evd isfix (List.combine fixnames fixdefs)
end
let interp_fixpoint l ntns =
@@ -1165,7 +1197,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind
List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
fixnames fixtypes fiximps in
let init_tac =
- Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
+ Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC)
fixdefs) in
let init_tac =
Option.map (List.map Proofview.V82.tactic) init_tac
@@ -1202,7 +1234,7 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n
List.map3 (fun id t (len,imps,_) -> ((id,pl),(t,(len,imps))))
fixnames fixtypes fiximps in
let init_tac =
- Some (List.map (Option.cata Tacmach.refine_no_check Tacticals.tclIDTAC)
+ Some (List.map (Option.cata (EConstr.of_constr %> Tacmach.refine_no_check) Tacticals.tclIDTAC)
fixdefs) in
let init_tac =
Option.map (List.map Proofview.V82.tactic) init_tac
@@ -1272,9 +1304,9 @@ let do_program_recursive local p fixkind fixl ntns =
let collect_evars id def typ imps =
(* Generalize by the recursive prototypes *)
let def =
- nf_evar evd (Termops.it_mkNamedLambda_or_LetIn def rec_sign)
+ EConstr.Unsafe.to_constr (nf_evar evd (Termops.it_mkNamedLambda_or_LetIn (EConstr.of_constr def) rec_sign))
and typ =
- nf_evar evd (Termops.it_mkNamedProd_or_LetIn typ rec_sign)
+ EConstr.Unsafe.to_constr (nf_evar evd (Termops.it_mkNamedProd_or_LetIn (EConstr.of_constr typ) rec_sign))
in
let evm = collect_evars_of_term evd def typ in
let evars, _, def, typ =
diff --git a/vernac/command.mli b/vernac/command.mli
index bccc22ae9..7cd0afeec 100644
--- a/vernac/command.mli
+++ b/vernac/command.mli
@@ -170,7 +170,7 @@ val do_cofixpoint :
(** Utils *)
-val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit
+val check_mutuality : Environ.env -> Evd.evar_map -> bool -> (Id.t * types) list -> unit
val declare_fix : ?opaque:bool -> definition_kind -> Universes.universe_binders -> Univ.universe_context -> Id.t ->
Safe_typing.private_constants Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference
diff --git a/vernac/discharge.ml b/vernac/discharge.ml
index e24d5e74f..b898f3e83 100644
--- a/vernac/discharge.ml
+++ b/vernac/discharge.ml
@@ -39,7 +39,7 @@ let detype_param =
let abstract_inductive hyps nparams inds =
let ntyp = List.length inds in
let nhyp = Context.Named.length hyps in
- let args = Context.Named.to_instance (List.rev hyps) in
+ let args = Context.Named.to_instance mkVar (List.rev hyps) in
let args = Array.of_list args in
let subs = List.init ntyp (fun k -> lift nhyp (mkApp(mkRel (k+1),args))) in
let inds' =
diff --git a/vernac/explainErr.ml b/vernac/explainErr.ml
index 2c99f35d4..04841c922 100644
--- a/vernac/explainErr.ml
+++ b/vernac/explainErr.ml
@@ -57,6 +57,7 @@ let process_vernac_interp_error exn = match fst exn with
mt() in
wrap_vernac_error exn (str "Universe inconsistency" ++ msg ++ str ".")
| TypeError(ctx,te) ->
+ let te = Himsg.map_ptype_error EConstr.of_constr te in
wrap_vernac_error exn (Himsg.explain_type_error ctx Evd.empty te)
| PretypeError(ctx,sigma,te) ->
wrap_vernac_error exn (Himsg.explain_pretype_error ctx sigma te)
diff --git a/vernac/himsg.ml b/vernac/himsg.ml
index 6cff805fc..17bb87f2a 100644
--- a/vernac/himsg.ml
+++ b/vernac/himsg.ml
@@ -29,50 +29,56 @@ module RelDecl = Context.Rel.Declaration
(* This simplifies the typing context of Cases clauses *)
(* hope it does not disturb other typing contexts *)
-let contract env lc =
+let contract env sigma lc =
+ let open EConstr in
let l = ref [] in
let contract_context decl env =
match decl with
- | LocalDef (_,c',_) when isRel c' ->
+ | LocalDef (_,c',_) when isRel sigma c' ->
l := (Vars.substl !l c') :: !l;
env
| _ ->
let t = Vars.substl !l (RelDecl.get_type decl) in
- let decl = decl |> RelDecl.map_name (named_hd env t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in
+ let decl = decl |> RelDecl.map_name (named_hd env sigma t) |> RelDecl.map_value (Vars.substl !l) |> RelDecl.set_type t in
l := (mkRel 1) :: List.map (Vars.lift 1) !l;
push_rel decl env
in
let env = process_rel_context contract_context env in
(env, List.map (Vars.substl !l) lc)
-let contract2 env a b = match contract env [a;b] with
+let contract2 env sigma a b = match contract env sigma [a;b] with
| env, [a;b] -> env,a,b | _ -> assert false
-let contract3 env a b c = match contract env [a;b;c] with
+let contract3 env sigma a b c = match contract env sigma [a;b;c] with
| env, [a;b;c] -> env,a,b,c | _ -> assert false
-let contract4 env a b c d = match contract env [a;b;c;d] with
+let contract4 env sigma a b c d = match contract env sigma [a;b;c;d] with
| env, [a;b;c;d] -> (env,a,b,c),d | _ -> assert false
-let contract1_vect env a v =
- match contract env (a :: Array.to_list v) with
+let contract1_vect env sigma a v =
+ match contract env sigma (a :: Array.to_list v) with
| env, a::l -> env,a,Array.of_list l
| _ -> assert false
-let rec contract3' env a b c = function
- | OccurCheck (evk,d) -> let x,d = contract4 env a b c d in x,OccurCheck(evk,d)
+let rec contract3' env sigma a b c = function
+ | OccurCheck (evk,d) ->
+ let x,d = contract4 env sigma a b c d in x,OccurCheck(evk, d)
| NotClean ((evk,args),env',d) ->
- let env',d,args = contract1_vect env' d args in
- contract3 env a b c,NotClean((evk,args),env',d)
+ let env',d,args = contract1_vect env' sigma d args in
+ contract3 env sigma a b c,NotClean((evk,args),env',d)
| ConversionFailed (env',t1,t2) ->
- let (env',t1,t2) = contract2 env' t1 t2 in
- contract3 env a b c, ConversionFailed (env',t1,t2)
+ let (env',t1,t2) = contract2 env' sigma t1 t2 in
+ contract3 env sigma a b c, ConversionFailed (env',t1,t2)
| NotSameArgSize | NotSameHead | NoCanonicalStructure
| MetaOccurInBody _ | InstanceNotSameType _ | ProblemBeyondCapabilities
- | UnifUnivInconsistency _ as x -> contract3 env a b c, x
+ | UnifUnivInconsistency _ as x -> contract3 env sigma a b c, x
| CannotSolveConstraint ((pb,env',t,u),x) ->
- let env',t,u = contract2 env' t u in
- let y,x = contract3' env a b c x in
+ let t = EConstr.of_constr t in
+ let u = EConstr.of_constr u in
+ let env',t,u = contract2 env' sigma t u in
+ let t = EConstr.Unsafe.to_constr t in
+ let u = EConstr.Unsafe.to_constr u in
+ let y,x = contract3' env sigma a b c x in
y,CannotSolveConstraint ((pb,env',t,u),x)
(** Ad-hoc reductions *)
@@ -82,12 +88,14 @@ let j_nf_betaiotaevar sigma j =
uj_type = Reductionops.nf_betaiota sigma j.uj_type }
let jv_nf_betaiotaevar sigma jl =
- Array.map (j_nf_betaiotaevar sigma) jl
+ Array.map (fun j -> j_nf_betaiotaevar sigma j) jl
(** Printers *)
let pr_lconstr c = quote (pr_lconstr c)
let pr_lconstr_env e s c = quote (pr_lconstr_env e s c)
+let pr_leconstr c = quote (pr_leconstr c)
+let pr_leconstr_env e s c = quote (pr_leconstr_env e s c)
let pr_ljudge_env e s c = let v,t = pr_ljudge_env e s c in (quote v,quote t)
(** A canonisation procedure for constr such that comparing there
@@ -144,7 +152,8 @@ let explicit_flags =
[print_implicits; print_coercions; print_no_symbol]; (** Then more! *)
[print_universes; print_implicits; print_coercions; print_no_symbol] (** and more! *) ]
-let pr_explicit env sigma t1 t2 = pr_explicit_aux env sigma t1 t2 explicit_flags
+let pr_explicit env sigma t1 t2 =
+ pr_explicit_aux env sigma (EConstr.to_constr sigma t1) (EConstr.to_constr sigma t2) explicit_flags
let pr_db env i =
try
@@ -177,9 +186,9 @@ let explain_bad_assumption env sigma j =
brk(1,1) ++ pc ++ spc () ++ str "of type" ++ spc () ++ pt ++ spc () ++
str "because this term is not a type."
-let explain_reference_variables id c =
+let explain_reference_variables sigma id c =
(* c is intended to be a global reference *)
- let pc = pr_global (Globnames.global_of_constr c) in
+ let pc = pr_global (fst (Termops.global_of_constr sigma c)) in
pc ++ strbrk " depends on the variable " ++ pr_id id ++
strbrk " which is not declared in the context."
@@ -196,9 +205,10 @@ let pr_puniverses f env (c,u) =
else mt())
let explain_elim_arity env sigma ind sorts c pj okinds =
- let env = make_all_name_different env in
+ let open EConstr in
+ let env = make_all_name_different env sigma in
let pi = pr_inductive env (fst ind) in
- let pc = pr_lconstr_env env sigma c in
+ let pc = pr_leconstr_env env sigma c in
let msg = match okinds with
| Some(kp,ki,explanation) ->
let pki = pr_sort_family ki in
@@ -211,7 +221,7 @@ let explain_elim_arity env sigma ind sorts c pj okinds =
| WrongArity ->
"wrong arity" in
let ppar = pr_disjunction (fun s -> quote (pr_sort_family s)) sorts in
- let ppt = pr_lconstr_env env sigma ((strip_prod_assum pj.uj_type)) in
+ let ppt = pr_leconstr_env env sigma (snd (decompose_prod_assum sigma pj.uj_type)) in
hov 0
(str "the return type has sort" ++ spc () ++ ppt ++ spc () ++
str "while it" ++ spc () ++ str "should be " ++ ppar ++ str ".") ++
@@ -231,10 +241,10 @@ let explain_elim_arity env sigma ind sorts c pj okinds =
let explain_case_not_inductive env sigma cj =
let cj = Evarutil.j_nf_evar sigma cj in
- let env = make_all_name_different env in
- let pc = pr_lconstr_env env sigma cj.uj_val in
- let pct = pr_lconstr_env env sigma cj.uj_type in
- match kind_of_term cj.uj_type with
+ let env = make_all_name_different env sigma in
+ let pc = pr_leconstr_env env sigma cj.uj_val in
+ let pct = pr_leconstr_env env sigma cj.uj_type in
+ match EConstr.kind sigma cj.uj_type with
| Evar _ ->
str "Cannot infer a type for this expression."
| _ ->
@@ -244,18 +254,17 @@ let explain_case_not_inductive env sigma cj =
let explain_number_branches env sigma cj expn =
let cj = Evarutil.j_nf_evar sigma cj in
- let env = make_all_name_different env in
- let pc = pr_lconstr_env env sigma cj.uj_val in
- let pct = pr_lconstr_env env sigma cj.uj_type in
+ let env = make_all_name_different env sigma in
+ let pc = pr_leconstr_env env sigma cj.uj_val in
+ let pct = pr_leconstr_env env sigma cj.uj_type in
str "Matching on term" ++ brk(1,1) ++ pc ++ spc () ++
str "of type" ++ brk(1,1) ++ pct ++ spc () ++
str "expects " ++ int expn ++ str " branches."
let explain_ill_formed_branch env sigma c ci actty expty =
- let simp t = Reduction.nf_betaiota env (Evarutil.nf_evar sigma t) in
- let c = Evarutil.nf_evar sigma c in
- let env = make_all_name_different env in
- let pc = pr_lconstr_env env sigma c in
+ let simp t = Reductionops.nf_betaiota sigma (Evarutil.nf_evar sigma t) in
+ let env = make_all_name_different env sigma in
+ let pc = pr_leconstr_env env sigma c in
let pa, pe = pr_explicit env sigma (simp actty) (simp expty) in
strbrk "In pattern-matching on term" ++ brk(1,1) ++ pc ++
spc () ++ strbrk "the branch for constructor" ++ spc () ++
@@ -265,7 +274,7 @@ let explain_ill_formed_branch env sigma c ci actty expty =
let explain_generalization env sigma (name,var) j =
let pe = pr_ne_context_of (str "In environment") env sigma in
- let pv = pr_ltype_env env sigma var in
+ let pv = pr_letype_env env sigma var in
let (pc,pt) = pr_ljudge_env (push_rel_assum (name,var) env) sigma j in
pe ++ str "Cannot generalize" ++ brk(1,1) ++ pv ++ spc () ++
str "over" ++ brk(1,1) ++ pc ++ str "," ++ spc () ++
@@ -277,27 +286,24 @@ let explain_unification_error env sigma p1 p2 = function
| Some e ->
let rec aux p1 p2 = function
| OccurCheck (evk,rhs) ->
- let rhs = Evarutil.nf_evar sigma rhs in
[str "cannot define " ++ quote (pr_existential_key sigma evk) ++
- strbrk " with term " ++ pr_lconstr_env env sigma rhs ++
+ strbrk " with term " ++ pr_leconstr_env env sigma rhs ++
strbrk " that would depend on itself"]
| NotClean ((evk,args),env,c) ->
- let c = Evarutil.nf_evar sigma c in
- let args = Array.map (Evarutil.nf_evar sigma) args in
[str "cannot instantiate " ++ quote (pr_existential_key sigma evk)
- ++ strbrk " because " ++ pr_lconstr_env env sigma c ++
+ ++ strbrk " because " ++ pr_leconstr_env env sigma c ++
strbrk " is not in its scope" ++
(if Array.is_empty args then mt() else
strbrk ": available arguments are " ++
- pr_sequence (pr_lconstr_env env sigma) (List.rev (Array.to_list args)))]
+ pr_sequence (pr_leconstr_env env sigma) (List.rev (Array.to_list args)))]
| NotSameArgSize | NotSameHead | NoCanonicalStructure ->
(* Error speaks from itself *) []
| ConversionFailed (env,t1,t2) ->
- if Term.eq_constr t1 p1 && Term.eq_constr t2 p2 then [] else
- let env = make_all_name_different env in
+ if EConstr.eq_constr sigma t1 p1 && EConstr.eq_constr sigma t2 p2 then [] else
+ let env = make_all_name_different env sigma in
let t1 = Evarutil.nf_evar sigma t1 in
let t2 = Evarutil.nf_evar sigma t2 in
- if not (Term.eq_constr t1 p1) || not (Term.eq_constr t2 p2) then
+ if not (EConstr.eq_constr sigma t1 p1) || not (EConstr.eq_constr sigma t2 p2) then
let t1, t2 = pr_explicit env sigma t1 t2 in
[str "cannot unify " ++ t1 ++ strbrk " and " ++ t2]
else []
@@ -318,11 +324,13 @@ let explain_unification_error env sigma p1 p2 = function
else
[str "universe inconsistency"]
| CannotSolveConstraint ((pb,env,t,u),e) ->
+ let t = EConstr.of_constr t in
+ let u = EConstr.of_constr u in
let t = Evarutil.nf_evar sigma t in
let u = Evarutil.nf_evar sigma u in
- let env = make_all_name_different env in
- (strbrk "cannot satisfy constraint " ++ pr_lconstr_env env sigma t ++
- str " == " ++ pr_lconstr_env env sigma u)
+ let env = make_all_name_different env sigma in
+ (strbrk "cannot satisfy constraint " ++ pr_leconstr_env env sigma t ++
+ str " == " ++ pr_leconstr_env env sigma u)
:: aux t u e
| ProblemBeyondCapabilities ->
[]
@@ -333,12 +341,12 @@ let explain_unification_error env sigma p1 p2 = function
prlist_with_sep pr_semicolon (fun x -> x) l ++ str ")"
let explain_actual_type env sigma j t reason =
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
let j = j_nf_betaiotaevar sigma j in
let t = Reductionops.nf_betaiota sigma t in
(** Actually print *)
let pe = pr_ne_context_of (str "In environment") env sigma in
- let pc = pr_lconstr_env env sigma (Environ.j_val j) in
+ let pc = pr_leconstr_env env sigma (Environ.j_val j) in
let (pt, pct) = pr_explicit env sigma t (Environ.j_type j) in
let ppreason = explain_unification_error env sigma j.uj_type t reason in
pe ++
@@ -353,7 +361,7 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
let exptyp = Evarutil.nf_evar sigma exptyp in
let actualtyp = Reductionops.nf_betaiota sigma actualtyp in
let rator = Evarutil.j_nf_evar sigma rator in
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
let actualtyp, exptyp = pr_explicit env sigma actualtyp exptyp in
let nargs = Array.length randl in
(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
@@ -379,15 +387,15 @@ let explain_cant_apply_bad_type env sigma (n,exptyp,actualtyp) rator randl =
let explain_cant_apply_not_functional env sigma rator randl =
let randl = Evarutil.jv_nf_evar sigma randl in
let rator = Evarutil.j_nf_evar sigma rator in
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
let nargs = Array.length randl in
(* let pe = pr_ne_context_of (str "in environment") env sigma in*)
- let pr = pr_lconstr_env env sigma rator.uj_val in
- let prt = pr_lconstr_env env sigma rator.uj_type in
+ let pr = pr_leconstr_env env sigma rator.uj_val in
+ let prt = pr_leconstr_env env sigma rator.uj_type in
let appl = prvect_with_sep fnl
(fun c ->
- let pc = pr_lconstr_env env sigma c.uj_val in
- let pct = pr_lconstr_env env sigma c.uj_type in
+ let pc = pr_leconstr_env env sigma c.uj_val in
+ let pct = pr_leconstr_env env sigma c.uj_type in
hov 2 (pc ++ spc () ++ str ": " ++ pct)) randl
in
str "Illegal application (Non-functional construction): " ++
@@ -405,7 +413,7 @@ let explain_unexpected_type env sigma actual_type expected_type =
str "where" ++ spc () ++ prexp ++ str " was expected."
let explain_not_product env sigma c =
- let c = Evarutil.nf_evar sigma c in
+ let c = EConstr.to_constr sigma c in
let pr = pr_lconstr_env env sigma c in
str "The type of this term is a product" ++ spc () ++
str "while it is expected to be" ++
@@ -414,6 +422,7 @@ let explain_not_product env sigma c =
(* TODO: use the names *)
(* (co)fixpoints *)
let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
+ let pr_lconstr_env env sigma c = pr_leconstr_env env sigma c in
let prt_name i =
match names.(i) with
Name id -> str "Recursive definition of " ++ pr_id id
@@ -428,7 +437,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
str "Recursive definition on" ++ spc () ++ pr_lconstr_env env sigma c ++
spc () ++ str "which should be an inductive type"
| RecursionOnIllegalTerm(j,(arg_env, arg),le,lt) ->
- let arg_env = make_all_name_different arg_env in
+ let arg_env = make_all_name_different arg_env sigma in
let called =
match names.(j) with
Name id -> pr_id id
@@ -494,7 +503,7 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
pr_ne_context_of (str "In environment") env sigma ++
st ++ str "." ++ fnl () ++
(try (* May fail with unresolved globals. *)
- let fixenv = make_all_name_different fixenv in
+ let fixenv = make_all_name_different fixenv sigma in
let pvd = pr_lconstr_env fixenv sigma vdefj.(i).uj_val in
str"Recursive definition is:" ++ spc () ++ pvd ++ str "."
with e when CErrors.noncritical e -> mt ())
@@ -502,8 +511,8 @@ let explain_ill_formed_rec_body env sigma err names i fixenv vdefj =
let explain_ill_typed_rec_body env sigma i names vdefj vargs =
let vdefj = Evarutil.jv_nf_evar sigma vdefj in
let vargs = Array.map (Evarutil.nf_evar sigma) vargs in
- let env = make_all_name_different env in
- let pvd = pr_lconstr_env env sigma vdefj.(i).uj_val in
+ let env = make_all_name_different env sigma in
+ let pvd = pr_leconstr_env env sigma vdefj.(i).uj_val in
let pvdt, pv = pr_explicit env sigma vdefj.(i).uj_type vargs.(i) in
str "The " ++
(match vdefj with [|_|] -> mt () | _ -> pr_nth (i+1) ++ spc ()) ++
@@ -512,16 +521,14 @@ let explain_ill_typed_rec_body env sigma i names vdefj vargs =
str "while it should be" ++ spc () ++ pv ++ str "."
let explain_cant_find_case_type env sigma c =
- let c = Evarutil.nf_evar sigma c in
- let env = make_all_name_different env in
- let pe = pr_lconstr_env env sigma c in
+ let env = make_all_name_different env sigma in
+ let pe = pr_leconstr_env env sigma c in
str "Cannot infer the return type of pattern-matching on" ++ ws 1 ++
pe ++ str "."
let explain_occur_check env sigma ev rhs =
- let rhs = Evarutil.nf_evar sigma rhs in
- let env = make_all_name_different env in
- let pt = pr_lconstr_env env sigma rhs in
+ let env = make_all_name_different env sigma in
+ let pt = pr_leconstr_env env sigma rhs in
str "Cannot define " ++ pr_existential_key sigma ev ++ str " with term" ++
brk(1,1) ++ pt ++ spc () ++ str "that would depend on itself."
@@ -567,16 +574,16 @@ let rec explain_evar_kind env sigma evk ty = function
| Evar_kinds.SubEvar evk' ->
let evi = Evd.find sigma evk' in
let pc = match evi.evar_body with
- | Evar_defined c -> pr_lconstr_env env sigma (Evarutil.nf_evar sigma c)
+ | Evar_defined c -> pr_leconstr_env env sigma (Evarutil.nf_evar sigma (EConstr.of_constr c))
| Evar_empty -> assert false in
- let ty' = Evarutil.nf_evar sigma evi.evar_concl in
+ let ty' = Evarutil.nf_evar sigma (EConstr.of_constr evi.evar_concl) in
pr_existential_key sigma evk ++ str " of type " ++ ty ++
str " in the partial instance " ++ pc ++
str " found for " ++ explain_evar_kind env sigma evk'
- (pr_lconstr_env env sigma ty') (snd evi.evar_source)
+ (pr_leconstr_env env sigma ty') (snd evi.evar_source)
let explain_typeclass_resolution env sigma evi k =
- match Typeclasses.class_of_constr evi.evar_concl with
+ match Typeclasses.class_of_constr sigma (EConstr.of_constr evi.evar_concl) with
| Some _ ->
let env = Evd.evar_filtered_env evi in
fnl () ++ str "Could not find an instance for " ++
@@ -589,7 +596,7 @@ let explain_placeholder_kind env sigma c e =
| Some (SeveralInstancesFound n) ->
strbrk " (several distinct possible type class instances found)"
| None ->
- match Typeclasses.class_of_constr c with
+ match Typeclasses.class_of_constr sigma (EConstr.of_constr c) with
| Some _ -> strbrk " (no type class instance found)"
| _ -> mt ()
@@ -619,7 +626,7 @@ let explain_wrong_case_info env (ind,u) ci =
spc () ++ pc ++ str "."
let explain_cannot_unify env sigma m n e =
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
let m = Evarutil.nf_evar sigma m in
let n = Evarutil.nf_evar sigma n in
let pm, pn = pr_explicit env sigma m n in
@@ -629,18 +636,19 @@ let explain_cannot_unify env sigma m n e =
str "with" ++ brk(1,1) ++ pn ++ ppreason ++ str "."
let explain_cannot_unify_local env sigma m n subn =
- let pm = pr_lconstr_env env sigma m in
- let pn = pr_lconstr_env env sigma n in
- let psubn = pr_lconstr_env env sigma subn in
+ let pm = pr_leconstr_env env sigma m in
+ let pn = pr_leconstr_env env sigma n in
+ let psubn = pr_leconstr_env env sigma subn in
str "Unable to unify" ++ brk(1,1) ++ pm ++ spc () ++
str "with" ++ brk(1,1) ++ pn ++ spc () ++ str "as" ++ brk(1,1) ++
psubn ++ str " contains local variables."
let explain_refiner_cannot_generalize env sigma ty =
str "Cannot find a well-typed generalisation of the goal with type: " ++
- pr_lconstr_env env sigma ty ++ str "."
+ pr_leconstr_env env sigma ty ++ str "."
let explain_no_occurrence_found env sigma c id =
+ let c = EConstr.to_constr sigma c in
str "Found no subterm matching " ++ pr_lconstr_env env sigma c ++
str " in " ++
(match id with
@@ -648,20 +656,24 @@ let explain_no_occurrence_found env sigma c id =
| None -> str"the current goal") ++ str "."
let explain_cannot_unify_binding_type env sigma m n =
- let pm = pr_lconstr_env env sigma m in
- let pn = pr_lconstr_env env sigma n in
+ let pm = pr_leconstr_env env sigma m in
+ let pn = pr_leconstr_env env sigma n in
str "This binding has type" ++ brk(1,1) ++ pm ++ spc () ++
str "which should be unifiable with" ++ brk(1,1) ++ pn ++ str "."
let explain_cannot_find_well_typed_abstraction env sigma p l e =
+ let p = EConstr.to_constr sigma p in
str "Abstracting over the " ++
str (String.plural (List.length l) "term") ++ spc () ++
- hov 0 (pr_enum (pr_lconstr_env env sigma) l) ++ spc () ++
+ hov 0 (pr_enum (fun c -> pr_lconstr_env env sigma (EConstr.to_constr sigma c)) l) ++ spc () ++
str "leads to a term" ++ spc () ++ pr_lconstr_goal_style_env env sigma p ++
spc () ++ str "which is ill-typed." ++
(match e with None -> mt () | Some e -> fnl () ++ str "Reason is: " ++ e)
let explain_wrong_abstraction_type env sigma na abs expected result =
+ let abs = EConstr.to_constr sigma abs in
+ let expected = EConstr.to_constr sigma expected in
+ let result = EConstr.to_constr sigma result in
let ppname = match na with Name id -> pr_id id ++ spc () | _ -> mt () in
str "Cannot instantiate metavariable " ++ ppname ++ strbrk "of type " ++
pr_lconstr_env env sigma expected ++ strbrk " with abstraction " ++
@@ -673,6 +685,7 @@ let explain_abstraction_over_meta _ m n =
pr_name m ++ spc () ++ str "and " ++ pr_name n ++ str "."
let explain_non_linear_unification env sigma m t =
+ let t = EConstr.to_constr sigma t in
strbrk "Cannot unambiguously instantiate " ++
pr_name m ++ str ":" ++
strbrk " which would require to abstract twice on " ++
@@ -680,11 +693,11 @@ let explain_non_linear_unification env sigma m t =
let explain_unsatisfied_constraints env sigma cst =
strbrk "Unsatisfied constraints: " ++
- Univ.pr_constraints (Evd.pr_evd_level sigma) cst ++
+ Univ.pr_constraints (Termops.pr_evd_level sigma) cst ++
spc () ++ str "(maybe a bugged tactic)."
let explain_type_error env sigma err =
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
match err with
| UnboundRel n ->
explain_unbound_rel env sigma n
@@ -695,7 +708,7 @@ let explain_type_error env sigma err =
| BadAssumption c ->
explain_bad_assumption env sigma c
| ReferenceVariables (id,c) ->
- explain_reference_variables id c
+ explain_reference_variables sigma id c
| ElimArity (ind, aritylst, c, pj, okinds) ->
explain_elim_arity env sigma ind aritylst c pj okinds
| CaseNotInductive cj ->
@@ -734,12 +747,16 @@ let explain_cannot_unify_occurrences env sigma nested ((cl2,pos2),t2) ((cl1,pos1
str "Found nested occurrences of the pattern at positions " ++
int pos1 ++ strbrk " and " ++ pr_position (cl2,pos2) ++ str "."
else
- let ppreason = match e with None -> mt() | Some (c1,c2,e) -> explain_unification_error env sigma c1 c2 (Some e) in
+ let ppreason = match e with
+ | None -> mt()
+ | Some (c1,c2,e) ->
+ explain_unification_error env sigma c1 c2 (Some e)
+ in
str "Found incompatible occurrences of the pattern" ++ str ":" ++
- spc () ++ str "Matched term " ++ pr_lconstr_env env sigma t2 ++
+ spc () ++ str "Matched term " ++ pr_lconstr_env env sigma (EConstr.to_constr sigma t2) ++
strbrk " at position " ++ pr_position (cl2,pos2) ++
strbrk " is not compatible with matched term " ++
- pr_lconstr_env env sigma t1 ++ strbrk " at position " ++
+ pr_lconstr_env env sigma (EConstr.to_constr sigma t1) ++ strbrk " at position " ++
pr_position (cl1,pos1) ++ ppreason ++ str "."
let pr_constraints printenv env sigma evars cstrs =
@@ -759,7 +776,7 @@ let pr_constraints printenv env sigma evars cstrs =
(fun (ev, evi) -> fnl () ++ pr_existential_key sigma ev ++
str " : " ++ pr_lconstr_env env' sigma evi.evar_concl ++ fnl ()) l
in
- h 0 (pe ++ evs ++ pr_evar_constraints cstrs)
+ h 0 (pe ++ evs ++ pr_evar_constraints sigma cstrs)
else
let filter evk _ = Evar.Map.mem evk evars in
pr_evar_map_filter ~with_univs:false filter sigma
@@ -795,23 +812,23 @@ let explain_unsatisfiable_constraints env sigma constr comp =
let explain_pretype_error env sigma err =
let env = Evardefine.env_nf_betaiotaevar sigma env in
- let env = make_all_name_different env in
+ let env = make_all_name_different env sigma in
match err with
| CantFindCaseType c -> explain_cant_find_case_type env sigma c
| ActualTypeNotCoercible (j,t,e) ->
let {uj_val = c; uj_type = actty} = j in
- let (env, c, actty, expty), e = contract3' env c actty t e in
+ let (env, c, actty, expty), e = contract3' env sigma c actty t e in
let j = {uj_val = c; uj_type = actty} in
explain_actual_type env sigma j expty (Some e)
| UnifOccurCheck (ev,rhs) -> explain_occur_check env sigma ev rhs
| UnsolvableImplicit (evk,exp) -> explain_unsolvable_implicit env sigma evk exp
| VarNotFound id -> explain_var_not_found env id
| UnexpectedType (actual,expect) ->
- let env, actual, expect = contract2 env actual expect in
+ let env, actual, expect = contract2 env sigma actual expect in
explain_unexpected_type env sigma actual expect
| NotProduct c -> explain_not_product env sigma c
| CannotUnify (m,n,e) ->
- let env, m, n = contract2 env m n in
+ let env, m, n = contract2 env sigma m n in
explain_cannot_unify env sigma m n e
| CannotUnifyLocal (m,n,sn) -> explain_cannot_unify_local env sigma m n sn
| CannotGeneralize ty -> explain_refiner_cannot_generalize env sigma ty
@@ -888,7 +905,7 @@ let explain_not_match_error = function
quote (Printer.safe_pr_lconstr_env env Evd.empty t2)
| IncompatibleConstraints cst ->
str " the expected (polymorphic) constraints do not imply " ++
- quote (Univ.pr_constraints (Evd.pr_evd_level Evd.empty) cst)
+ quote (Univ.pr_constraints (Termops.pr_evd_level Evd.empty) cst)
let explain_signature_mismatch l spec why =
str "Signature components for label " ++ pr_label l ++
@@ -1004,6 +1021,7 @@ let explain_module_internalization_error = function
(* Typeclass errors *)
let explain_not_a_class env c =
+ let c = EConstr.to_constr Evd.empty c in
pr_constr_env env Evd.empty c ++ str" is not a declared type class."
let explain_unbound_method env cid id =
@@ -1059,7 +1077,7 @@ let explain_non_linear_proof c =
spc () ++ str "because a metavariable has several occurrences."
let explain_meta_in_type c =
- str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_lconstr c ++
+ str "In refiner, a meta appears in the type " ++ brk(1,1) ++ pr_leconstr c ++
str " of another meta"
let explain_no_such_hyp id =
@@ -1092,7 +1110,7 @@ let error_ill_formed_inductive env c v =
let error_ill_formed_constructor env id c v nparams nargs =
let pv = pr_lconstr_env env Evd.empty v in
- let atomic = Int.equal (nb_prod c) 0 in
+ let atomic = Int.equal (nb_prod Evd.empty (EConstr.of_constr c)) (** FIXME *) 0 in
str "The type of constructor" ++ brk(1,1) ++ pr_id id ++ brk(1,1) ++
str "is not valid;" ++ brk(1,1) ++
strbrk (if atomic then "it must be " else "its conclusion must be ") ++
@@ -1111,7 +1129,7 @@ let error_ill_formed_constructor env id c v nparams nargs =
let pr_ltype_using_barendregt_convention_env env c =
(* Use goal_concl_style as an approximation of Barendregt's convention (?) *)
- quote (pr_goal_concl_style_env env Evd.empty c)
+ quote (pr_goal_concl_style_env env Evd.empty (EConstr.of_constr c))
let error_bad_ind_parameters env c n v1 v2 =
let pc = pr_ltype_using_barendregt_convention_env env c in
@@ -1193,7 +1211,8 @@ let explain_recursion_scheme_error = function
(* Pattern-matching errors *)
let explain_bad_pattern env sigma cstr ty =
- let env = make_all_name_different env in
+ let ty = EConstr.to_constr sigma ty in
+ let env = make_all_name_different env sigma in
let pt = pr_lconstr_env env sigma ty in
let pc = pr_constructor env cstr in
str "Found the constructor " ++ pc ++ brk(1,1) ++
@@ -1236,13 +1255,15 @@ let explain_non_exhaustive env pats =
spc () ++ hov 0 (prlist_with_sep pr_comma pr_cases_pattern pats)
let explain_cannot_infer_predicate env sigma typs =
- let env = make_all_name_different env in
+ let inj c = EConstr.to_constr sigma c in
+ let typs = Array.map_to_list (fun (c1, c2) -> (inj c1, inj c2)) typs in
+ let env = make_all_name_different env sigma in
let pr_branch (cstr,typ) =
let cstr,_ = decompose_app cstr in
str "For " ++ pr_lconstr_env env sigma cstr ++ str ": " ++ pr_lconstr_env env sigma typ
in
str "Unable to unify the types found in the branches:" ++
- spc () ++ hov 0 (prlist_with_sep fnl pr_branch (Array.to_list typs))
+ spc () ++ hov 0 (prlist_with_sep fnl pr_branch typs)
let explain_pattern_matching_error env sigma = function
| BadPattern (c,t) ->
@@ -1260,8 +1281,48 @@ let explain_pattern_matching_error env sigma = function
| CannotInferPredicate typs ->
explain_cannot_infer_predicate env sigma typs
+let map_pguard_error f = function
+| NotEnoughAbstractionInFixBody -> NotEnoughAbstractionInFixBody
+| RecursionNotOnInductiveType c -> RecursionNotOnInductiveType (f c)
+| RecursionOnIllegalTerm (n, (env, c), l1, l2) -> RecursionOnIllegalTerm (n, (env, f c), l1, l2)
+| NotEnoughArgumentsForFixCall n -> NotEnoughArgumentsForFixCall n
+| CodomainNotInductiveType c -> CodomainNotInductiveType (f c)
+| NestedRecursiveOccurrences -> NestedRecursiveOccurrences
+| UnguardedRecursiveCall c -> UnguardedRecursiveCall (f c)
+| RecCallInTypeOfAbstraction c -> RecCallInTypeOfAbstraction (f c)
+| RecCallInNonRecArgOfConstructor c -> RecCallInNonRecArgOfConstructor (f c)
+| RecCallInTypeOfDef c -> RecCallInTypeOfDef (f c)
+| RecCallInCaseFun c -> RecCallInCaseFun (f c)
+| RecCallInCaseArg c -> RecCallInCaseArg (f c)
+| RecCallInCasePred c -> RecCallInCasePred (f c)
+| NotGuardedForm c -> NotGuardedForm (f c)
+| ReturnPredicateNotCoInductive c -> ReturnPredicateNotCoInductive (f c)
+
+let map_ptype_error f = function
+| UnboundRel n -> UnboundRel n
+| UnboundVar id -> UnboundVar id
+| NotAType j -> NotAType (on_judgment f j)
+| BadAssumption j -> BadAssumption (on_judgment f j)
+| ReferenceVariables (id, c) -> ReferenceVariables (id, f c)
+| ElimArity (pi, dl, c, j, ar) -> ElimArity (pi, dl, f c, on_judgment f j, ar)
+| CaseNotInductive j -> CaseNotInductive (on_judgment f j)
+| WrongCaseInfo (pi, ci) -> WrongCaseInfo (pi, ci)
+| NumberBranches (j, n) -> NumberBranches (on_judgment f j, n)
+| IllFormedBranch (c, pc, t1, t2) -> IllFormedBranch (f c, pc, f t1, f t2)
+| Generalization ((na, t), j) -> Generalization ((na, f t), on_judgment f j)
+| ActualType (j, t) -> ActualType (on_judgment f j, f t)
+| CantApplyBadType ((n, c1, c2), j, vj) ->
+ CantApplyBadType ((n, f c1, f c2), on_judgment f j, Array.map (on_judgment f) vj)
+| CantApplyNonFunctional (j, jv) -> CantApplyNonFunctional (on_judgment f j, Array.map (on_judgment f) jv)
+| IllFormedRecBody (ge, na, n, env, jv) ->
+ IllFormedRecBody (map_pguard_error f ge, na, n, env, Array.map (on_judgment f) jv)
+| IllTypedRecBody (n, na, jv, t) ->
+ IllTypedRecBody (n, na, Array.map (on_judgment f) jv, Array.map f t)
+| UnsatisfiedConstraints g -> UnsatisfiedConstraints g
+
let explain_reduction_tactic_error = function
| Tacred.InvalidAbstraction (env,sigma,c,(env',e)) ->
+ let e = map_ptype_error EConstr.of_constr e in
str "The abstracted term" ++ spc () ++
quote (pr_goal_concl_style_env env sigma c) ++
spc () ++ str "is not well typed." ++ fnl () ++
diff --git a/vernac/himsg.mli b/vernac/himsg.mli
index ced54fd27..6915ea921 100644
--- a/vernac/himsg.mli
+++ b/vernac/himsg.mli
@@ -40,3 +40,6 @@ val explain_module_error : Modops.module_typing_error -> std_ppcmds
val explain_module_internalization_error :
Modintern.module_internalization_error -> std_ppcmds
+
+val map_pguard_error : ('c -> 'd) -> 'c pguard_error -> 'd pguard_error
+val map_ptype_error : ('c -> 'd) -> ('c, 'c) ptype_error -> ('d, 'd) ptype_error
diff --git a/vernac/indschemes.ml b/vernac/indschemes.ml
index f7e3f0d95..9ba4e46e4 100644
--- a/vernac/indschemes.ml
+++ b/vernac/indschemes.ml
@@ -319,7 +319,7 @@ let warn_cannot_build_congruence =
strbrk "Cannot build congruence scheme because eq is not found")
let declare_congr_scheme ind =
- if Hipattern.is_equality_type (mkInd ind) then begin
+ if Hipattern.is_equality_type Evd.empty (EConstr.of_constr (mkInd ind)) (** FIXME *) then begin
if
try Coqlib.check_required_library Coqlib.logic_module_name; true
with e when CErrors.noncritical e -> false
@@ -409,7 +409,8 @@ let do_mutual_induction_scheme lnamedepindsort =
in
let sigma, listdecl = Indrec.build_mutual_induction_scheme env0 sigma lrecspec in
let declare decl fi lrecref =
- let decltype = Retyping.get_type_of env0 sigma decl in
+ let decltype = Retyping.get_type_of env0 sigma (EConstr.of_constr decl) in
+ let decltype = EConstr.to_constr sigma decltype in
let proof_output = Future.from_val ((decl,Univ.ContextSet.empty),Safe_typing.empty_private_constants) in
let cst = define fi UserIndividualRequest sigma proof_output (Some decltype) in
ConstRef cst :: lrecref
@@ -478,7 +479,7 @@ let build_combined_scheme env schemes =
let (c, t) = List.hd defs in
let ctx, ind, nargs = find_inductive t in
(* Number of clauses, including the predicates quantification *)
- let prods = nb_prod t - (nargs + 1) in
+ let prods = nb_prod Evd.empty (EConstr.of_constr t) - (nargs + 1) (** FIXME *) in
let coqand = Coqlib.build_coq_and () and coqconj = Coqlib.build_coq_conj () in
let relargs = rel_vect 0 prods in
let concls = List.rev_map
@@ -494,7 +495,7 @@ let build_combined_scheme env schemes =
let ctx, _ =
list_split_rev_at prods
(List.rev_map (fun (x, y) -> LocalAssum (x, y)) ctx) in
- let typ = it_mkProd_wo_LetIn concl_typ ctx in
+ let typ = List.fold_left (fun d c -> Term.mkProd_wo_LetIn c d) concl_typ ctx in
let body = it_mkLambda_or_LetIn concl_bod ctx in
(body, typ)
diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml
index 409676276..993a2c260 100644
--- a/vernac/lemmas.ml
+++ b/vernac/lemmas.ml
@@ -108,7 +108,7 @@ let find_mutually_recursive_statements thms =
(* Cofixpoint or fixpoint w/o explicit decreasing argument *)
| None | Some (None, CStructRec) ->
let whnf_hyp_hds = map_rel_context_in_env
- (fun env c -> fst (whd_all_stack env Evd.empty c))
+ (fun env c -> EConstr.Unsafe.to_constr (fst (whd_all_stack env Evd.empty (EConstr.of_constr c))))
(Global.env()) hyps in
let ind_hyps =
List.flatten (List.map_i (fun i decl ->
@@ -122,8 +122,8 @@ let find_mutually_recursive_statements thms =
[]) 0 (List.rev whnf_hyp_hds)) in
let ind_ccl =
let cclenv = push_rel_context hyps (Global.env()) in
- let whnf_ccl,_ = whd_all_stack cclenv Evd.empty ccl in
- match kind_of_term whnf_ccl with
+ 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
| 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 ->
@@ -393,7 +393,7 @@ let start_proof_univs id ?pl kind sigma ?terminator ?sign c ?init_tac ?(compute_
let rec_tac_initializer finite guard thms snl =
if finite then
- match List.map (fun ((id,_),(t,_)) -> (id,t)) thms with
+ match List.map (fun ((id,_),(t,_)) -> (id,EConstr.of_constr t)) thms with
| (id,_)::l -> Tactics.mutual_cofix id l 0
| _ -> assert false
else
@@ -401,7 +401,7 @@ let rec_tac_initializer finite guard thms snl =
let nl = match snl with
| None -> List.map succ (List.map List.last guard)
| Some nl -> nl
- in match List.map2 (fun ((id,_),(t,_)) n -> (id,n,t)) thms nl with
+ in match List.map2 (fun ((id,_),(t,_)) n -> (id,n, EConstr.of_constr t)) thms nl with
| (id,n,_)::l -> Tactics.mutual_fix id n l 0
| _ -> assert false
@@ -449,7 +449,7 @@ let start_proof_with_initialization kind ctx recguard thms snl hook =
List.iter (fun (strength,ref,imps) ->
maybe_declare_manual_implicits false ref imps;
call_hook (fun exn -> exn) hook strength ref) thms_data in
- start_proof_univs id ?pl kind ctx t ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
+ start_proof_univs id ?pl kind ctx (EConstr.of_constr t) ?init_tac (fun ctx -> mk_hook (hook ctx)) ~compute_guard:guard
let start_proof_com ?inference_hook kind thms hook =
let env0 = Global.env () in
@@ -466,7 +466,7 @@ let start_proof_com ?inference_hook kind thms hook =
evdref := solve_remaining_evars flags env !evdref Evd.empty;
let ids = List.map RelDecl.get_name ctx in
(compute_proof_name (pi1 kind) sopt,
- (nf_evar !evdref (it_mkProd_or_LetIn t' ctx),
+ (EConstr.Unsafe.to_constr (nf_evar !evdref (EConstr.it_mkProd_or_LetIn t' ctx)),
(ids, imps @ lift_implicits (List.length ids) imps'),
guard)))
thms in
@@ -518,6 +518,7 @@ let save_proof ?proof = function
| None ->
let pftree = Pfedit.get_pftreestate () in
let id, k, typ = Pfedit.current_proof_statement () in
+ let typ = EConstr.Unsafe.to_constr typ in
let universes = Proof.initial_euctx pftree in
(* This will warn if the proof is complete *)
let pproofs, _univs =
diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli
index 39c089be9..681413a29 100644
--- a/vernac/lemmas.mli
+++ b/vernac/lemmas.mli
@@ -19,17 +19,17 @@ val call_hook :
Future.fix_exn -> 'a declaration_hook -> Decl_kinds.locality -> Globnames.global_reference -> 'a
(** A hook start_proof calls on the type of the definition being started *)
-val set_start_hook : (types -> unit) -> unit
+val set_start_hook : (EConstr.types -> unit) -> unit
val start_proof : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map ->
?terminator:(lemma_possible_guards -> unit declaration_hook -> Proof_global.proof_terminator) ->
- ?sign:Environ.named_context_val -> types ->
+ ?sign:Environ.named_context_val -> EConstr.types ->
?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards ->
unit declaration_hook -> unit
val start_proof_univs : Id.t -> ?pl:universe_binders -> goal_kind -> Evd.evar_map ->
?terminator:(lemma_possible_guards -> (Evd.evar_universe_context option -> unit declaration_hook) -> Proof_global.proof_terminator) ->
- ?sign:Environ.named_context_val -> types ->
+ ?sign:Environ.named_context_val -> EConstr.types ->
?init_tac:unit Proofview.tactic -> ?compute_guard:lemma_possible_guards ->
(Evd.evar_universe_context option -> unit declaration_hook) -> unit
diff --git a/vernac/obligations.ml b/vernac/obligations.ml
index 6f3921903..ea2401b5c 100644
--- a/vernac/obligations.ml
+++ b/vernac/obligations.ml
@@ -144,10 +144,11 @@ let trunc_named_context n ctx =
List.firstn (len - n) ctx
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
- | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (Termops.pop b) else None
+ | Prod (_, _, b) -> if noccurn 1 b then chop_product (pred n) (pop b) else None
| _ -> None
let evar_dependencies evm oev =
@@ -266,7 +267,7 @@ let pperror cmd = CErrors.user_err ~hdr:"Program" cmd
let error s = pperror (str s)
let reduce c =
- Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty c
+ EConstr.Unsafe.to_constr (Reductionops.clos_norm_flags CClosure.betaiota (Global.env ()) Evd.empty (EConstr.of_constr c))
exception NoObligations of Id.t option
@@ -398,7 +399,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 (Termops.strip_outer_cast t) with
+ match kind_of_term (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
| _ ->
@@ -523,7 +524,7 @@ let compute_possible_guardness_evidences (n,_) fixbody fixtype =
but doing it properly involves delta-reduction, and it finally
doesn't seem to worth the effort (except for huge mutual
fixpoints ?) *)
- let m = Termops.nb_prod fixtype in
+ let m = Termops.nb_prod Evd.empty (EConstr.of_constr fixtype) (** FIXME *) in
let ctx = fst (decompose_prod_n_assum m fixtype) in
List.map_i (fun i _ -> i) 0 ctx
@@ -536,8 +537,10 @@ let declare_mutual_definition l =
List.split3
(List.map (fun x ->
let subs, typ = (subst_body true x) in
- let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len subs) in
- let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len typ) in
+ let term = snd (Reductionops.splay_lam_n (Global.env ()) Evd.empty len (EConstr.of_constr subs)) in
+ let typ = snd (Reductionops.splay_prod_n (Global.env ()) Evd.empty len (EConstr.of_constr typ)) in
+ let term = EConstr.Unsafe.to_constr term in
+ let typ = EConstr.Unsafe.to_constr typ in
x.prg_reduce term, x.prg_reduce typ, x.prg_implicits) l)
in
(* let fixdefs = List.map reduce_fix fixdefs in *)
@@ -817,7 +820,7 @@ let rec string_of_list sep f = function
let solve_by_tac name evi t poly ctx =
let id = name in
- let concl = evi.evar_concl in
+ let concl = EConstr.of_constr evi.evar_concl in
(* spiwack: the status is dropped. *)
let (entry,_,ctx') = Pfedit.build_constant_by_tactic
id ~goal_kind:(goal_kind poly) ctx evi.evar_hyps concl (Tacticals.New.tclCOMPLETE t) in
@@ -935,7 +938,7 @@ let rec solve_obligation prg num tac =
Proof_global.make_terminator
(obligation_terminator prg.prg_name num guard hook auto) in
let hook ctx = Lemmas.mk_hook (obligation_hook prg obl num auto ctx) in
- let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd obl.obl_type ~terminator hook in
+ let () = Lemmas.start_proof_univs ~sign:prg.prg_sign obl.obl_name kind evd (EConstr.of_constr obl.obl_type) ~terminator hook in
let _ = Pfedit.by !default_tactic in
Option.iter (fun tac -> Pfedit.set_end_tac tac) tac
diff --git a/vernac/record.ml b/vernac/record.ml
index 288d3391b..8b4b7606f 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -66,7 +66,8 @@ let interp_fields_evars env evars impls_env nots l =
List.fold_left2
(fun (env, uimpls, params, impls) no ((loc, i), b, t) ->
let t', impl = interp_type_evars_impls env evars ~impls t in
- let b' = Option.map (fun x -> fst (interp_casted_constr_evars_impls env evars ~impls x t')) b in
+ let b' = Option.map (fun x -> EConstr.Unsafe.to_constr (fst (interp_casted_constr_evars_impls env evars ~impls x t'))) b in
+ let t' = EConstr.Unsafe.to_constr t' in
let impls =
match i with
| Anonymous -> impls
@@ -84,7 +85,7 @@ let compute_constructor_level evars env l =
List.fold_right (fun d (env, univ) ->
let univ =
if is_local_assum d then
- let s = Retyping.get_sort_of env evars (RelDecl.get_type d) in
+ let s = Retyping.get_sort_of env evars (EConstr.of_constr (RelDecl.get_type d)) in
Univ.sup (univ_of_sort s) univ
else univ
in (push_rel d env, univ))
@@ -114,6 +115,7 @@ let typecheck_params_and_fields def id pl t ps nots fs =
Loc.raise ~loc (Stream.Error "pattern with quote not allowed in record parameters.")) ps
in
let impls_env, ((env1,newps), imps) = interp_context_evars env0 evars ps in
+ let newps = List.map (fun d -> Termops.map_rel_decl EConstr.Unsafe.to_constr d) newps in
let t', template = match t with
| Some t ->
let env = push_rel_context newps env0 in
@@ -122,6 +124,8 @@ let typecheck_params_and_fields def id pl t ps nots fs =
| CSort (_, Misctypes.GType []) -> true | _ -> false in
let s = interp_type_evars env evars ~impls:empty_internalization_env t in
let sred = Reductionops.whd_all env !evars s in
+ let s = EConstr.Unsafe.to_constr s in
+ let sred = EConstr.Unsafe.to_constr sred in
(match kind_of_term sred with
| Sort s' ->
(if poly then
@@ -164,7 +168,7 @@ let typecheck_params_and_fields def id pl t ps nots fs =
let evars, nf = Evarutil.nf_evars_and_universes evars in
let newps = Context.Rel.map nf newps in
let newfs = Context.Rel.map nf newfs in
- let ce t = Pretyping.check_evars env0 Evd.empty evars t in
+ let ce t = Pretyping.check_evars env0 Evd.empty evars (EConstr.of_constr t) in
List.iter (iter_constr ce) (List.rev newps);
List.iter (iter_constr ce) (List.rev newfs);
Evd.universe_context ?names:pl evars, nf arity, template, imps, newps, impls, newfs
@@ -271,8 +275,8 @@ let declare_projections indsp ?(kind=StructureComponent) binder_name coers field
let ctx = Univ.instantiate_univ_context mib.mind_universes in
let indu = indsp, u in
let r = mkIndU (indsp,u) in
- let rp = applist (r, Context.Rel.to_extended_list 0 paramdecls) in
- let paramargs = Context.Rel.to_extended_list 1 paramdecls in (*def in [[params;x:rp]]*)
+ let rp = applist (r, Context.Rel.to_extended_list mkRel 0 paramdecls) in
+ let paramargs = Context.Rel.to_extended_list mkRel 1 paramdecls in (*def in [[params;x:rp]]*)
let x = Name binder_name in
let fields = instantiate_possibly_recursive_type indu paramdecls fields in
let lifted_fields = Termops.lift_rel_context 1 fields in
@@ -365,17 +369,17 @@ let structure_signature ctx =
| [decl] ->
let env = Environ.empty_named_context_val in
let evm = Sigma.Unsafe.of_evar_map evm in
- let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in
+ let Sigma (_, evm, _) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in
let evm = Sigma.to_evar_map evm in
evm
| decl::tl ->
let env = Environ.empty_named_context_val in
let evm = Sigma.Unsafe.of_evar_map evm in
- let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (RelDecl.get_type decl) in
+ let Sigma (ev, evm, _) = Evarutil.new_pure_evar env evm (EConstr.of_constr (RelDecl.get_type decl)) in
let evm = Sigma.to_evar_map evm in
let new_tl = Util.List.map_i
(fun pos decl ->
- RelDecl.map_type (fun t -> Termops.replace_term (mkRel pos) (mkEvar(ev,[||])) t) decl) 1 tl in
+ RelDecl.map_type (fun t -> EConstr.Unsafe.to_constr (Termops.replace_term evm (EConstr.mkRel pos) (EConstr.mkEvar(ev,[||])) (EConstr.of_constr t))) decl) 1 tl in
deps_to_evar evm new_tl in
deps_to_evar Evd.empty (List.rev ctx)
@@ -384,7 +388,7 @@ open Typeclasses
let declare_structure finite poly ctx id idbuild paramimpls params arity template
fieldimpls fields ?(kind=StructureComponent) ?name is_coe coers sign =
let nparams = List.length params and nfields = List.length fields in
- let args = Context.Rel.to_extended_list nfields params in
+ let args = Context.Rel.to_extended_list mkRel nfields params in
let ind = applist (mkRel (1+nparams+nfields), args) in
let type_constructor = it_mkProd_or_LetIn ind fields in
let binder_name =
@@ -485,7 +489,7 @@ let declare_class finite def poly ctx id idbuild paramimpls params arity
in
let ctx_context =
List.map (fun decl ->
- match Typeclasses.class_of_constr (RelDecl.get_type decl) with
+ match Typeclasses.class_of_constr Evd.empty (EConstr.of_constr (RelDecl.get_type decl)) with
| Some (_, ((cl,_), _)) -> Some (cl.cl_impl, true)
| None -> None)
params, params
diff --git a/vernac/search.ml b/vernac/search.ml
index 540573843..6279b17ae 100644
--- a/vernac/search.ml
+++ b/vernac/search.ml
@@ -177,20 +177,20 @@ let prioritize_search seq fn =
(** This function tries to see whether the conclusion matches a pattern. *)
(** FIXME: this is quite dummy, we may find a more efficient algorithm. *)
-let rec pattern_filter pat ref env typ =
- let typ = Termops.strip_outer_cast typ in
- if Constr_matching.is_matching env Evd.empty pat typ then true
- else match kind_of_term typ with
+let rec pattern_filter pat ref env sigma typ =
+ let typ = Termops.strip_outer_cast sigma typ in
+ if Constr_matching.is_matching env sigma pat typ then true
+ else match EConstr.kind sigma typ with
| Prod (_, _, typ)
- | LetIn (_, _, _, typ) -> pattern_filter pat ref env typ
+ | LetIn (_, _, _, typ) -> pattern_filter pat ref env sigma typ
| _ -> false
-let rec head_filter pat ref env typ =
- let typ = Termops.strip_outer_cast typ in
- if Constr_matching.is_matching_head env Evd.empty pat typ then true
- else match kind_of_term typ with
+let rec head_filter pat ref env sigma typ =
+ let typ = Termops.strip_outer_cast sigma typ in
+ if Constr_matching.is_matching_head env sigma pat typ then true
+ else match EConstr.kind sigma typ with
| Prod (_, _, typ)
- | LetIn (_, _, _, typ) -> head_filter pat ref env typ
+ | LetIn (_, _, _, typ) -> head_filter pat ref env sigma typ
| _ -> false
let full_name_of_reference ref =
@@ -217,7 +217,7 @@ let name_of_reference ref = Id.to_string (basename_of_global ref)
let search_about_filter query gr env typ = match query with
| GlobSearchSubPattern pat ->
- Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat typ
+ Constr_matching.is_matching_appsubterm ~closed:false env Evd.empty pat (EConstr.of_constr typ)
| GlobSearchString s ->
String.string_contains ~where:(name_of_reference gr) ~what:s
@@ -228,7 +228,7 @@ let search_pattern gopt pat mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
module_filter mods ref env typ &&
- pattern_filter pat ref env typ &&
+ pattern_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -252,8 +252,8 @@ let search_rewrite gopt pat mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
module_filter mods ref env typ &&
- (pattern_filter pat1 ref env typ ||
- pattern_filter pat2 ref env typ) &&
+ (pattern_filter pat1 ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) ||
+ pattern_filter pat2 ref env Evd.empty (EConstr.of_constr typ)) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -267,7 +267,7 @@ let search_by_head gopt pat mods pr_search =
let blacklist_filter = blacklist_filter_aux () in
let filter ref env typ =
module_filter mods ref env typ &&
- head_filter pat ref env typ &&
+ head_filter pat ref env Evd.empty (* FIXME *) (EConstr.of_constr typ) &&
blacklist_filter ref env typ
in
let iter ref env typ =
@@ -331,12 +331,12 @@ let interface_search =
toggle (Str.string_match regexp id 0) flag
in
let match_type (pat, flag) =
- toggle (Constr_matching.is_matching env Evd.empty pat constr) flag
+ toggle (Constr_matching.is_matching env Evd.empty pat (EConstr.of_constr constr)) flag
in
let match_subtype (pat, flag) =
toggle
(Constr_matching.is_matching_appsubterm ~closed:false
- env Evd.empty pat constr) flag
+ env Evd.empty pat (EConstr.of_constr constr)) flag
in
let match_module (mdl, flag) =
toggle (Libnames.is_dirpath_prefix_of mdl path) flag
diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml
index ca03ba3f3..53d49ddbc 100644
--- a/vernac/vernacentries.ml
+++ b/vernac/vernacentries.ml
@@ -60,7 +60,7 @@ let show_proof () =
(* spiwack: this would probably be cooler with a bit of polishing. *)
let p = Proof_global.give_me_the_proof () in
let pprf = Proof.partial_proof p in
- Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_constr pprf)
+ Feedback.msg_notice (Pp.prlist_with_sep Pp.fnl Printer.pr_econstr pprf)
let show_node () =
(* spiwack: I'm have little clue what this function used to do. I deactivated it,
@@ -81,8 +81,8 @@ let show_universes () =
let gls = Proof.V82.subgoals pfts in
let sigma = gls.Evd.sigma in
let ctx = Evd.universe_context_set (Evd.nf_constraints sigma) in
- Feedback.msg_notice (Evd.pr_evar_universe_context (Evd.evar_universe_context sigma));
- Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Evd.pr_evd_level sigma) ctx)
+ Feedback.msg_notice (Termops.pr_evar_universe_context (Evd.evar_universe_context sigma));
+ Feedback.msg_notice (str"Normalized constraints: " ++ Univ.pr_universe_context_set (Termops.pr_evd_level sigma) ctx)
let show_prooftree () =
(* Spiwack: proof tree is currently not working *)
@@ -103,11 +103,12 @@ let try_print_subgoals () =
(* Simulate the Intro(s) tactic *)
let show_intro all =
+ let open EConstr in
let pf = get_pftreestate() in
let {Evd.it=gls ; sigma=sigma; } = Proof.V82.subgoals pf in
if not (List.is_empty gls) then begin
let gl = {Evd.it=List.hd gls ; sigma = sigma; } in
- let l,_= decompose_prod_assum (Termops.strip_outer_cast (pf_concl gl)) in
+ let l,_= decompose_prod_assum sigma (Termops.strip_outer_cast sigma (pf_concl gl)) in
if all then
let lid = Tactics.find_intro_names l gl in
Feedback.msg_notice (hov 0 (prlist_with_sep spc pr_id lid))
@@ -457,11 +458,12 @@ let start_proof_and_print k l hook =
let env = Evd.evar_filtered_env evi in
try
let concl = Evarutil.nf_evars_universes sigma evi.Evd.evar_concl in
+ let concl = EConstr.of_constr concl in
if Evarutil.has_undefined_evars sigma concl then raise Exit;
let c, _, ctx =
Pfedit.build_by_tactic env (Evd.evar_universe_context sigma)
concl (Tacticals.New.tclCOMPLETE tac)
- in Evd.set_universe_context sigma ctx, c
+ in Evd.set_universe_context sigma ctx, EConstr.of_constr c
with Logic_monad.TacticFailure e when Logic.catchable_exception e ->
error "The statement obligations could not be resolved \
automatically, write a statement definition first."
@@ -869,6 +871,7 @@ let vernac_set_used_variables e =
let env = Global.env () in
let tys =
List.map snd (Proof.initial_goals (Proof_global.give_me_the_proof ())) in
+ let tys = List.map EConstr.Unsafe.to_constr tys in
let l = Proof_using.process_expr env e tys in
let vars = Environ.named_context env in
List.iter (fun id ->
@@ -1230,7 +1233,7 @@ let vernac_reserve bl =
let env = Global.env() in
let sigma = Evd.from_env env in
let t,ctx = Constrintern.interp_type env sigma c in
- let t = Detyping.detype false [] env (Evd.from_ctx ctx) t in
+ let t = Detyping.detype false [] env (Evd.from_ctx ctx) (EConstr.of_constr t) in
let t,_ = Notation_ops.notation_constr_of_glob_constr (default_env ()) t in
Reserve.declare_reserved_type idl t)
in List.iter sb_decl bl
@@ -1567,6 +1570,7 @@ let get_current_context_of_args = function
let vernac_check_may_eval redexp glopt rc =
let (sigma, env) = get_current_context_of_args glopt in
let sigma', c = interp_open_constr env sigma rc in
+ let c = EConstr.Unsafe.to_constr c in
let sigma' = Evarconv.solve_unif_constraints_with_heuristics env sigma' in
Evarconv.check_problems_are_solved env sigma';
let sigma',nf = Evarutil.nf_evars_and_universes sigma' in
@@ -1574,14 +1578,16 @@ let vernac_check_may_eval redexp glopt rc =
let env = Environ.push_context uctx (Evarutil.nf_env_evar sigma' env) in
let c = nf c in
let j =
- if Evarutil.has_undefined_evars sigma' c then
- Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' c)
+ if Evarutil.has_undefined_evars sigma' (EConstr.of_constr c) then
+ Evarutil.j_nf_evar sigma' (Retyping.get_judgment_of env sigma' (EConstr.of_constr c))
else
(* OK to call kernel which does not support evars *)
- Arguments_renaming.rename_typing env c in
+ Termops.on_judgment EConstr.of_constr (Arguments_renaming.rename_typing env c)
+ in
match redexp with
| None ->
- let l = Evar.Set.union (Evd.evars_of_term j.Environ.uj_val) (Evd.evars_of_term j.Environ.uj_type) in
+ let evars_of_term c = Evarutil.undefined_evars_of_term sigma' c in
+ let l = Evar.Set.union (evars_of_term j.Environ.uj_val) (evars_of_term j.Environ.uj_type) in
let j = { j with Environ.uj_type = Reductionops.nf_betaiota sigma' j.Environ.uj_type } in
Feedback.msg_notice (print_judgment env sigma' j ++
pr_ne_evar_set (fnl () ++ str "where" ++ fnl ()) (mt ()) sigma' l ++
@@ -1641,7 +1647,7 @@ let print_about_hyp_globs ref_or_by_not glnumopt =
let natureofid = match decl with
| LocalAssum _ -> "Hypothesis"
| LocalDef (_,bdy,_) ->"Constant (let in)" in
- v 0 (pr_id id ++ str":" ++ pr_constr (NamedDecl.get_type decl) ++ fnl() ++ fnl()
+ v 0 (pr_id id ++ str":" ++ pr_econstr (NamedDecl.get_type decl) ++ fnl() ++ fnl()
++ str natureofid ++ str " of the goal context.")
with (* fallback to globals *)
| NoHyp | Not_found -> print_about ref_or_by_not
@@ -1893,7 +1899,7 @@ let vernac_check_guard () =
try
let { Evd.it=gl ; sigma=sigma } = Proof.V82.top_goal pts in
Inductiveops.control_only_guard (Goal.V82.env sigma gl)
- pfterm;
+ (EConstr.Unsafe.to_constr pfterm);
(str "The condition holds up to here")
with UserError(_,s) ->
(str ("Condition violated: ") ++s)