aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins
diff options
context:
space:
mode:
Diffstat (limited to 'plugins')
-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
61 files changed, 1747 insertions, 1369 deletions
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 8fba92e79..db8f3e4b2 100644
--- a/plugins/nsatz/nsatz.ml
+++ b/plugins/nsatz/nsatz.ml
@@ -570,6 +570,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 *)