From b4b90c5d2e8c413e1981c456c933f35679386f09 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Sat, 26 Nov 2016 16:18:47 +0100 Subject: Definining EConstr-based contexts. This removes quite a few unsafe casts. Unluckily, I had to reintroduce the old non-module based names for these data structures, because I could not reproduce easily the same hierarchy in EConstr. --- plugins/cc/cctac.ml | 8 ++----- plugins/decl_mode/decl_proof_instr.ml | 15 +++++++------ plugins/extraction/extraction.ml | 2 +- plugins/firstorder/instances.ml | 3 +-- plugins/firstorder/rules.ml | 2 +- plugins/fourier/fourierR.ml | 4 +++- plugins/funind/functional_principles_proofs.ml | 7 +++--- plugins/funind/functional_principles_types.ml | 16 ++++++++------ plugins/funind/indfun.ml | 2 +- plugins/funind/invfun.ml | 20 ++++++----------- plugins/funind/merge.ml | 2 +- plugins/funind/recdef.ml | 2 +- plugins/omega/coq_omega.ml | 30 +++++++++++--------------- plugins/rtauto/refl_tauto.ml | 28 +++++++++++++----------- plugins/rtauto/refl_tauto.mli | 6 +++--- plugins/ssrmatching/ssrmatching.ml4 | 4 ++-- 16 files changed, 72 insertions(+), 79 deletions(-) (limited to 'plugins') diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index 53c450116..5d894c677 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -28,10 +28,6 @@ open Proofview.Notations module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - RelDecl.LocalAssum (na, inj t) - let reference dir s = lazy (Coqlib.gen_reference "CC" dir s) let _f_equal = reference ["Init";"Logic"] "f_equal" @@ -160,7 +156,7 @@ let rec quantified_atom_of_constr env sigma nrels term = let patts=patterns_of_constr env sigma nrels atom in `Nrule patts else - quantified_atom_of_constr (Environ.push_rel (local_assum (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 @@ -175,7 +171,7 @@ let litteral_of_constr env sigma term= else begin try - quantified_atom_of_constr (Environ.push_rel (local_assum (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 diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index da971fffb..adc4ad8a3 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -49,7 +49,7 @@ let clear ids { it = goal; sigma } = user_err (str "Cannot clear " ++ pr_id id) in let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in { it = [gl]; sigma } @@ -638,7 +638,7 @@ let assume_tac hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label)) + Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) st.st_label)) hyps tclIDTAC gls let assume_hyps_or_theses hyps gls = @@ -648,7 +648,7 @@ let assume_hyps_or_theses hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,c)))) nam) + Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr c)))) nam) | Hprop {st_label=nam;st_it=Thesis (tk)} -> tclTHEN (push_intro_tac @@ -660,7 +660,7 @@ let assume_st hyps gls = (fun st -> tclTHEN (push_intro_tac - (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) st.st_label)) + (fun id -> Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) st.st_label)) hyps tclIDTAC gls let assume_st_letin hyps gls = @@ -669,7 +669,7 @@ let assume_st_letin hyps gls = tclTHEN (push_intro_tac (fun id -> - Proofview.V82.of_tactic (convert_hyp (LocalDef (id, fst st.st_it, snd st.st_it)))) st.st_label)) + Proofview.V82.of_tactic (convert_hyp (LocalDef (id, EConstr.of_constr (fst st.st_it), EConstr.of_constr (snd st.st_it))))) st.st_label)) hyps tclIDTAC gls (* suffices *) @@ -763,7 +763,7 @@ let rec consider_match may_intro introduced available expected gls = error "Not enough sub-hypotheses to match statements." (* should tell which ones *) | id::rest_ids,(Hvar st | Hprop st)::rest -> - tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,st.st_it)))) + tclIFTHENELSE (Proofview.V82.of_tactic (convert_hyp (LocalAssum (id,EConstr.of_constr st.st_it)))) begin match st.st_label with Anonymous -> @@ -834,13 +834,13 @@ let define_tac id args body gls = (* tactics for reconsider *) let cast_tac id_or_thesis typ gls = + let typ = EConstr.of_constr typ in match id_or_thesis with | This id -> Proofview.V82.of_tactic (id |> pf_get_hyp gls |> NamedDecl.set_id id |> NamedDecl.set_type typ |> convert_hyp) gls | Thesis (For _ ) -> error "\"thesis for ...\" is not applicable here." | Thesis Plain -> - let typ = EConstr.of_constr typ in Proofview.V82.of_tactic (convert_concl typ DEFAULTcast) gls (* per cases *) @@ -1290,6 +1290,7 @@ let rec execute_cases fix_name per_info tacnext args objs nhrec tree gls = let f_ids typ = let sign = (prod_assum (Term.prod_applist typ params)) in + let sign = List.map (fun d -> Termops.map_rel_decl EConstr.of_constr d) sign in find_intro_names sign gls in let constr_args_ids = Array.map f_ids gen_arities in let case_term = diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 7b7e746f2..8744eacd3 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -858,7 +858,7 @@ and extract_fix env mle i (fi,ti,ci as recd) mlt = let decomp_lams_eta_n n m env c t = 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,c)) rels 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. *) diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index ef8172de4..9dc2a51a6 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -117,10 +117,9 @@ let mk_open_instance id idc gl m t= let nid=(fresh_id avoid var_id gl) in let evmap = Sigma.Unsafe.of_evar_map evmap in let Sigma ((c, _), evmap, _) = Evarutil.new_type_evar env evmap Evd.univ_flexible in - let c = EConstr.Unsafe.to_constr c 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 diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 36bd91ab6..a60fd4d8f 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -42,7 +42,7 @@ let wrap n b continue seq 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 (EConstr.Unsafe.to_constr (pf_concl gls)) seq1 gls else seq1 in diff --git a/plugins/fourier/fourierR.ml b/plugins/fourier/fourierR.ml index ec73fccb5..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 @@ -504,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)) diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 91b17b9a4..bc64b079c 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -236,7 +236,7 @@ let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta -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_leconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_leconstr t ); @@ -315,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!"); - (pop end_of_type,ctxt_size,mkLetIn (RelDecl.get_name decl, witness, EConstr.of_constr (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) ) @@ -544,7 +544,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = (scan_type new_context new_t') with Failure "NoChange" -> (* Last thing todo : push the rel in the context and continue *) - scan_type (local_assum (x,t_x) :: context) t' + scan_type (LocalAssum (x,t_x) :: context) t' end end else @@ -933,6 +933,7 @@ 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 (project g) hyp) keep diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index d0d44b34b..e845db3bc 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -33,9 +33,9 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = 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 -> @@ -56,7 +56,7 @@ 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 @@ -87,17 +87,19 @@ 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 mkProd_or_LetIn princ_type_info.indarg - (EConstr.Unsafe.to_constr princ_type_info.concl) + princ_type_info.concl ) princ_type_info.args ) 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 @@ -240,7 +242,7 @@ 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) @@ -251,7 +253,7 @@ let change_property_sort evd toSort princ princName = 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); Term.compose_prod args (mkSort toSort) @@ -270,7 +272,7 @@ 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 *) diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml index 1cde4420e..a7489fb7b 100644 --- a/plugins/funind/indfun.ml +++ b/plugins/funind/indfun.ml @@ -19,7 +19,7 @@ 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 sigma (EConstr.of_constr (RelDecl.get_type decl)))) 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 diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index dcec2cb74..8f1420940 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -26,12 +26,6 @@ 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) - (* Some pretty printing function for debugging purpose *) let pr_binding prc = @@ -147,7 +141,7 @@ let generate_type evd g_to_f f graph i = let fun_ctxt,res_type = match ctxt with | [] | [_] -> anomaly (Pp.str "Not a valid context") - | decl :: fun_ctxt -> fun_ctxt, EConstr.of_constr (RelDecl.get_type decl) + | decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl in let rec args_from_decl i accu = function | [] -> accu @@ -187,12 +181,12 @@ let generate_type evd g_to_f f graph i = \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \] i*) let pre_ctxt = - local_assum (Name res_id, lift 1 res_type) :: local_def (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt + LocalAssum (Name res_id, lift 1 res_type) :: LocalDef (Name fv_id, mkApp (f,args_as_rels), res_type) :: fun_ctxt in (*i and we can return the solution depending on which lemma type we are defining i*) if g_to_f - then local_assum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph - else local_assum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph + then LocalAssum (Anonymous,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args),graph + else LocalAssum (Anonymous,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied),graph (* @@ -280,7 +274,7 @@ 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 evd (EConstr.of_constr (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 @@ -477,7 +471,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) (project g) x (EConstr.of_constr 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) @@ -695,7 +689,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 (project g) (EConstr.of_constr (RelDecl.get_type decl)))) + (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (RelDecl.get_type decl))) ) branches in diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index 2840193a9..691385fad 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -976,7 +976,7 @@ let funify_branches relinfo nfuns branch = | Rel i -> let reali = i-shift in (reali>=0 && reali false in (* FIXME: *) - LocalDef (Anonymous,mkProp,mkProp) + LocalDef (Anonymous,EConstr.mkProp,EConstr.mkProp) let relprinctype_to_funprinctype relprinctype nfuns = diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 56c6ab054..f5ea32878 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -693,7 +693,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 (project g) expr (EConstr.of_constr (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 diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 9e0d591b6..8c2f0f53f 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -1706,10 +1706,6 @@ let onClearedName2 id tac = Tacticals.New.tclTHENLIST [ introduction id1; introduction id2; tac id1 id2 ] end }) -let local_assum (na, t) = - let inj = EConstr.Unsafe.to_constr in - LocalAssum (na, inj t) - let rec is_Prop sigma c = match EConstr.kind sigma c with | Sort (Prop Null) -> true | Cast (c,_,_) -> is_Prop sigma c @@ -1725,24 +1721,24 @@ let destructure_hyps = | decl::lit -> let i = NamedDecl.get_id decl in Proofview.tclEVARMAP >>= fun sigma -> - begin try match destructurate_prop sigma (EConstr.of_constr (NamedDecl.get_type decl)) with + 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]) -> (Tacticals.New.tclTHENS (elim_id i) - [ onClearedName i (fun i -> (loop (local_assum (i,t1)::lit))); - onClearedName i (fun i -> (loop (local_assum (i,t2)::lit))) ]) + [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit))); + onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (local_assum (i1,t1) :: local_assum (i2,t2) :: lit))) + loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> Tacticals.New.tclTHEN (elim_id i) (onClearedName2 i (fun i1 i2 -> - loop (local_assum (i1,mkArrow t1 t2) :: local_assum (i2,mkArrow t2 t1) :: lit))) + loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit))) | Kimp(t1,t2) -> (* t1 and t2 might be in Type rather than Prop. For t1, the decidability check will ensure being Prop. *) @@ -1753,7 +1749,7 @@ let destructure_hyps = (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_or (mk_not t1) t2) :: lit)))) + (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit)))) ] else loop lit @@ -1764,7 +1760,7 @@ let destructure_hyps = (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(And,[t1;t2]) -> let d1 = decidability t1 in @@ -1773,7 +1769,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_and, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit)))) ] | Kapp(Iff,[t1;t2]) -> let d1 = decidability t1 in @@ -1783,7 +1779,7 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_iff, [| t1; t2; d1; d2; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i, mk_or (mk_and t1 (mk_not t2)) + (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2)) (mk_and (mk_not t1) t2)) :: lit)))) ] | Kimp(t1,t2) -> @@ -1795,14 +1791,14 @@ let destructure_hyps = [mkApp (Lazy.force coq_not_imp, [| t1; t2; d1; mkVar i |])]); (onClearedName i (fun i -> - (loop (local_assum (i,mk_and t1 (mk_not t2)) :: lit)))) + (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit)))) ] | Kapp(Not,[t]) -> let d = decidability t in Tacticals.New.tclTHENLIST [ (generalize_tac [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); - (onClearedName i (fun i -> (loop (local_assum (i,t) :: lit)))) + (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit)))) ] | Kapp(op,[t1;t2]) -> (try @@ -1835,12 +1831,12 @@ let destructure_hyps = match destructurate_type sigma (pf_nf typ) with | Kapp(Nat,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_neq, [| t1;t2|]))) + (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) decl)) (loop lit)) | Kapp(Z,_) -> (Tacticals.New.tclTHEN - (convert_hyp_no_check (NamedDecl.set_type (EConstr.Unsafe.to_constr (mkApp (Lazy.force coq_Zne, [| t1;t2|]))) + (convert_hyp_no_check (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) decl)) (loop lit)) | _ -> loop lit diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index 475380512..adb11f4f8 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -66,19 +66,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 @@ -90,13 +89,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 EConstr.Vars.noccurn (Tacmach.project gls) 1 (EConstr.of_constr b) && + if noccurn sigma 1 b && Retyping.get_sort_family_of - (pf_env gls) (Tacmach.project gls) (EConstr.of_constr 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 @@ -113,7 +116,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 @@ -134,9 +137,9 @@ 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 (fun c -> Termops.local_occur_var Evd.empty (** FIXME *) id (EConstr.of_constr c)) 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) (EConstr.of_constr typ) != InProp) + (pf_env gls) (Tacmach.project gls) typ != InProp) then hrec else @@ -264,7 +267,6 @@ let rtauto_tac gls= if Retyping.get_sort_family_of (pf_env gls) (Tacmach.project gls) gl != InProp then user_err ~hdr:"rtauto" (Pp.str "goal should be in Prop") in - let gl = EConstr.Unsafe.to_constr gl in let glf=make_form gamma gls gl in let hyps=make_hyps gamma gls [gl] (pf_hyps gls) in let formula= 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/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index bf3e2ac1c..d4579c3a1 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -838,7 +838,7 @@ 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 + 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 @@ -1091,7 +1091,7 @@ let thin id sigma goal = | None -> sigma | Some (hyps, concl) -> let sigma = !evdref in - let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps (EConstr.of_constr concl) (Goal.V82.extra sigma goal) in + let (gl,ev,sigma) = Goal.V82.mk_goal sigma hyps concl (Goal.V82.extra sigma goal) in let sigma = Goal.V82.partial_solution_to sigma goal gl ev in sigma -- cgit v1.2.3