aboutsummaryrefslogtreecommitdiffhomepage
path: root/tactics
diff options
context:
space:
mode:
authorGravatar Maxime Dénès <mail@maximedenes.fr>2017-04-11 00:28:47 +0200
committerGravatar Maxime Dénès <mail@maximedenes.fr>2017-04-11 00:28:47 +0200
commit835be3a05e28eb6e26f703a034f22b2c6c61acaa (patch)
tree00ecf04840ba027c3c71f8503d9811c8a5dc1d2e /tactics
parent0980dbb1740c8d48d8ff0c516929f27f8cea854d (diff)
parent2e6a89238dc7197057d0da80a16f4b4b1e41bfd8 (diff)
Merge PR#379: Introducing evar-insensitive constrs
Diffstat (limited to 'tactics')
-rw-r--r--tactics/auto.ml70
-rw-r--r--tactics/auto.mli1
-rw-r--r--tactics/autorewrite.ml20
-rw-r--r--tactics/btermdn.ml33
-rw-r--r--tactics/btermdn.mli2
-rw-r--r--tactics/class_tactics.ml158
-rw-r--r--tactics/class_tactics.mli5
-rw-r--r--tactics/contradiction.ml27
-rw-r--r--tactics/contradiction.mli1
-rw-r--r--tactics/eauto.ml69
-rw-r--r--tactics/eauto.mli1
-rw-r--r--tactics/elim.ml18
-rw-r--r--tactics/elim.mli1
-rw-r--r--tactics/eqdecide.ml36
-rw-r--r--tactics/eqdecide.mli2
-rw-r--r--tactics/eqschemes.ml85
-rw-r--r--tactics/equality.ml218
-rw-r--r--tactics/equality.mli1
-rw-r--r--tactics/hints.ml188
-rw-r--r--tactics/hints.mli15
-rw-r--r--tactics/hipattern.ml249
-rw-r--r--tactics/hipattern.mli34
-rw-r--r--tactics/inv.ml62
-rw-r--r--tactics/inv.mli1
-rw-r--r--tactics/leminv.ml40
-rw-r--r--tactics/leminv.mli1
-rw-r--r--tactics/tacticals.ml161
-rw-r--r--tactics/tacticals.mli42
-rw-r--r--tactics/tactics.ml688
-rw-r--r--tactics/tactics.mli25
-rw-r--r--tactics/term_dnet.ml7
31 files changed, 1223 insertions, 1038 deletions
diff --git a/tactics/auto.ml b/tactics/auto.ml
index 7558a707e..74cb7a364 100644
--- a/tactics/auto.ml
+++ b/tactics/auto.ml
@@ -6,14 +6,15 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*
-*)
+module CVars = Vars
+
open Pp
open Util
open CErrors
open Names
open Vars
open Termops
+open EConstr
open Environ
open Tacmach
open Genredexpr
@@ -83,25 +84,26 @@ let connect_hint_clenv poly (c, _, ctx) clenv gl =
if poly then
(** Refresh the instance of the hint *)
let (subst, ctx) = Universes.fresh_universe_context_set_instance ctx in
- let map c = Vars.subst_univs_level_constr subst c in
+ let map c = CVars.subst_univs_level_constr subst c in
+ let emap c = Vars.subst_univs_level_constr subst c in
let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in
(** Only metas are mentioning the old universes. *)
let clenv = {
- templval = Evd.map_fl map clenv.templval;
- templtyp = Evd.map_fl map clenv.templtyp;
+ templval = Evd.map_fl emap clenv.templval;
+ templtyp = Evd.map_fl emap clenv.templtyp;
evd = Evd.map_metas map evd;
env = Proofview.Goal.env gl;
} in
- clenv, map c
+ clenv, emap c
else
let evd = Evd.merge_context_set Evd.univ_flexible evd ctx in
{ clenv with evd = evd ; env = Proofview.Goal.env gl }, c
in clenv, c
let unify_resolve poly flags ((c : raw_hint), clenv) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let clenv, c = connect_hint_clenv poly c clenv gl in
- let clenv = Tacmach.New.of_old (fun gl -> clenv_unique_resolver ~flags clenv gl) gl in
+ let clenv = clenv_unique_resolver ~flags clenv gl in
Clenvtac.clenv_refine false clenv
end }
@@ -298,13 +300,13 @@ let flags_of_state st =
let auto_flags_of_state st =
auto_unif_flags_of full_transparent_state st false
-let hintmap_of secvars hdc concl =
+let hintmap_of sigma secvars hdc concl =
match hdc with
| None -> Hint_db.map_none ~secvars
| Some hdc ->
- if occur_existential concl then
- Hint_db.map_existential ~secvars hdc concl
- else Hint_db.map_auto ~secvars hdc concl
+ if occur_existential sigma concl then
+ Hint_db.map_existential sigma ~secvars hdc concl
+ else Hint_db.map_auto sigma ~secvars hdc concl
let exists_evaluable_reference env = function
| EvalConstRef _ -> true
@@ -328,25 +330,26 @@ let rec trivial_fail_db dbg mod_delta db_list local_db =
end })
in
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let sigma = Tacmach.New.project gl in
let secvars = compute_secvars gl in
Tacticals.New.tclFIRST
((dbg_assumption dbg)::intro_tac::
(List.map Tacticals.New.tclCOMPLETE
- (trivial_resolve dbg mod_delta db_list local_db secvars concl)))
+ (trivial_resolve sigma dbg mod_delta db_list local_db secvars concl)))
end }
-and my_find_search_nodelta db_list local_db secvars hdc concl =
+and my_find_search_nodelta sigma db_list local_db secvars hdc concl =
List.map (fun hint -> (None,hint))
- (List.map_append (hintmap_of secvars hdc concl) (local_db::db_list))
+ (List.map_append (hintmap_of sigma secvars hdc concl) (local_db::db_list))
and my_find_search mod_delta =
if mod_delta then my_find_search_delta
else my_find_search_nodelta
-and my_find_search_delta db_list local_db secvars hdc concl =
- let f = hintmap_of secvars hdc concl in
- if occur_existential concl then
+and my_find_search_delta sigma db_list local_db secvars hdc concl =
+ let f = hintmap_of sigma secvars hdc concl in
+ if occur_existential sigma concl then
List.map_append
(fun db ->
if Hint_db.use_dn db then
@@ -368,8 +371,8 @@ and my_find_search_delta db_list local_db secvars hdc concl =
match hdc with None -> Hint_db.map_none ~secvars db
| Some hdc ->
if (Id.Pred.is_empty ids && Cpred.is_empty csts)
- then Hint_db.map_auto ~secvars hdc concl db
- else Hint_db.map_existential ~secvars hdc concl db
+ then Hint_db.map_auto sigma ~secvars hdc concl db
+ else Hint_db.map_existential sigma ~secvars hdc concl db
in auto_flags_of_state st, l
in List.map (fun x -> (Some flags,x)) l)
(local_db::db_list)
@@ -402,23 +405,23 @@ and tac_of_hint dbg db_list local_db concl (flags, ({pat=p; code=t;poly=poly;db=
in
tclLOG dbg pr_hint (run_hint t tactic)
-and trivial_resolve dbg mod_delta db_list local_db secvars cl =
+and trivial_resolve sigma dbg mod_delta db_list local_db secvars cl =
try
let head =
- try let hdconstr = decompose_app_bound cl in
+ try let hdconstr = decompose_app_bound sigma cl in
Some hdconstr
with Bound -> None
in
List.map (tac_of_hint dbg db_list local_db cl)
(priority
- (my_find_search mod_delta db_list local_db secvars head cl))
+ (my_find_search mod_delta sigma db_list local_db secvars head cl))
with Not_found -> []
(** The use of the "core" database can be de-activated by passing
"nocore" amongst the databases. *)
let trivial ?(debug=Off) lems dbnames =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = make_db_list dbnames in
@@ -429,7 +432,7 @@ let trivial ?(debug=Off) lems dbnames =
end }
let full_trivial ?(debug=Off) lems =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = current_pure_db () in
@@ -449,15 +452,15 @@ let h_trivial ?(debug=Off) lems l = gen_trivial ~debug lems l
(* The classical Auto tactic *)
(**************************************************************************)
-let possible_resolve dbg mod_delta db_list local_db secvars cl =
+let possible_resolve sigma dbg mod_delta db_list local_db secvars cl =
try
let head =
- try let hdconstr = decompose_app_bound cl in
+ try let hdconstr = decompose_app_bound sigma cl in
Some hdconstr
with Bound -> None
in
List.map (tac_of_hint dbg db_list local_db cl)
- (my_find_search mod_delta db_list local_db secvars head cl)
+ (my_find_search mod_delta sigma db_list local_db secvars head cl)
with Not_found -> []
let extend_local_db decl db gl =
@@ -487,13 +490,14 @@ let search d n mod_delta db_list local_db =
Tacticals.New.tclORELSE0 (dbg_assumption d)
(Tacticals.New.tclORELSE0 (intro_register d (search d n) local_db)
( Proofview.Goal.enter { enter = begin fun gl ->
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let sigma = Tacmach.New.project gl in
let secvars = compute_secvars gl in
let d' = incr_dbg d in
Tacticals.New.tclFIRST
(List.map
(fun ntac -> Tacticals.New.tclTHEN ntac (search d' (n-1) local_db))
- (possible_resolve d mod_delta db_list local_db secvars concl))
+ (possible_resolve sigma d mod_delta db_list local_db secvars concl))
end }))
end []
in
@@ -502,7 +506,7 @@ let search d n mod_delta db_list local_db =
let default_search_depth = ref 5
let delta_auto debug mod_delta n lems dbnames =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = make_db_list dbnames in
@@ -525,7 +529,7 @@ let new_auto ?(debug=Off) n = delta_auto debug true n
let default_auto = auto !default_search_depth [] []
let delta_full_auto ?(debug=Off) mod_delta n lems =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let db_list = current_pure_db () in
diff --git a/tactics/auto.mli b/tactics/auto.mli
index de0dbd483..32710e347 100644
--- a/tactics/auto.mli
+++ b/tactics/auto.mli
@@ -10,6 +10,7 @@
open Names
open Term
+open EConstr
open Clenv
open Pattern
open Decl_kinds
diff --git a/tactics/autorewrite.ml b/tactics/autorewrite.ml
index dae1cc9f1..e58ec5a31 100644
--- a/tactics/autorewrite.ml
+++ b/tactics/autorewrite.ml
@@ -92,7 +92,7 @@ type raw_rew_rule = Loc.t * constr Univ.in_universe_context_set * bool * Genarg.
let one_base general_rewrite_maybe_in tac_main bas =
let lrul = find_rewrites bas in
let try_rewrite dir ctx c tc =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let subst, ctx' = Universes.fresh_universe_context_set_instance ctx in
let c' = Vars.subst_univs_level_constr subst c in
@@ -122,7 +122,7 @@ let autorewrite ?(conds=Naive) tac_main lbas =
Tacticals.New.tclTHEN tac
(one_base (fun dir c tac ->
let tac = (tac, conds) in
- general_rewrite dir AllOccurrences true false ~tac c)
+ general_rewrite dir AllOccurrences true false ~tac (EConstr.of_constr c))
tac_main bas))
(Proofview.tclUNIT()) lbas))
@@ -165,7 +165,7 @@ let autorewrite_multi_in ?(conds=Naive) idl tac_main lbas =
| _ -> assert false) (* there must be at least an hypothesis *)
| _ -> assert false (* rewriting cannot complete a proof *)
in
- let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y z w) in
+ let general_rewrite_in x y z w = Proofview.V82.tactic (general_rewrite_in x y (EConstr.of_constr z) w) in
Tacticals.New.tclMAP (fun id ->
Tacticals.New.tclREPEAT_MAIN (Proofview.tclPROGRESS
(List.fold_left (fun tac bas ->
@@ -257,12 +257,12 @@ type hypinfo = {
let decompose_applied_relation metas env sigma c ctype left2right =
let find_rel ty =
- let eqclause = Clenv.mk_clenv_from_env env sigma None (c,ty) in
+ let eqclause = Clenv.mk_clenv_from_env env sigma None (EConstr.of_constr c,ty) in
let eqclause =
if metas then eqclause
else clenv_pose_metas_as_evars eqclause (Evd.undefined_metas eqclause.evd)
in
- let (equiv, args) = decompose_app (Clenv.clenv_type eqclause) in
+ let (equiv, args) = decompose_app (EConstr.Unsafe.to_constr (Clenv.clenv_type eqclause)) in
let rec split_last_two = function
| [c1;c2] -> [],(c1, c2)
| x::y::z ->
@@ -272,11 +272,13 @@ let decompose_applied_relation metas env sigma c ctype left2right =
try
let others,(c1,c2) = split_last_two args in
let ty1, ty2 =
- Typing.unsafe_type_of env eqclause.evd c1, Typing.unsafe_type_of env eqclause.evd c2
+ Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c1), Typing.unsafe_type_of env eqclause.evd (EConstr.of_constr c2)
in
+ let ty = EConstr.Unsafe.to_constr ty in
+ let ty1 = EConstr.Unsafe.to_constr ty1 in
(* if not (evd_convertible env eqclause.evd ty1 ty2) then None *)
(* else *)
- Some { hyp_cl=eqclause; hyp_prf=(Clenv.clenv_value eqclause); hyp_ty = ty;
+ Some { hyp_cl=eqclause; hyp_prf=EConstr.Unsafe.to_constr (Clenv.clenv_value eqclause); hyp_ty = ty;
hyp_car=ty1; hyp_rel=mkApp (equiv, Array.of_list others);
hyp_l2r=left2right; hyp_left=c1; hyp_right=c2; }
with Not_found -> None
@@ -290,12 +292,12 @@ let decompose_applied_relation metas env sigma c ctype left2right =
| None -> None
let find_applied_relation metas loc env sigma c left2right =
- let ctype = Typing.unsafe_type_of env sigma c in
+ let ctype = Typing.unsafe_type_of env sigma (EConstr.of_constr c) in
match decompose_applied_relation metas env sigma c ctype left2right with
| Some c -> c
| None ->
user_err ~loc ~hdr:"decompose_applied_relation"
- (str"The type" ++ spc () ++ Printer.pr_constr_env env sigma ctype ++
+ (str"The type" ++ spc () ++ Printer.pr_econstr_env env sigma ctype ++
spc () ++ str"of this term does not end with an applied relation.")
(* To add rewriting rules to a base *)
diff --git a/tactics/btermdn.ml b/tactics/btermdn.ml
index 491bc8b4a..b4a235ba8 100644
--- a/tactics/btermdn.ml
+++ b/tactics/btermdn.ml
@@ -8,6 +8,7 @@
open Util
open Term
+open EConstr
open Names
open Pattern
open Globnames
@@ -38,18 +39,18 @@ let decomp_pat =
in
decrec []
-let decomp =
- let rec decrec acc c = match kind_of_term c with
+let decomp sigma t =
+ let rec decrec acc c = match EConstr.kind sigma c with
| App (f,l) -> decrec (Array.fold_right (fun a l -> a::l) l acc) f
| Proj (p, c) -> (mkConst (Projection.constant p), c :: acc)
| Cast (c1,_,_) -> decrec acc c1
| _ -> (c,acc)
in
- decrec []
+ decrec [] t
-let constr_val_discr t =
- let c, l = decomp t in
- match kind_of_term c with
+let constr_val_discr sigma t =
+ let c, l = decomp sigma t in
+ match EConstr.kind sigma c with
| Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l)
| Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l)
| Var id -> Label(GRLabel (VarRef id),l)
@@ -66,9 +67,9 @@ let constr_pat_discr t =
| PRef ((VarRef v) as ref), args -> Some(GRLabel ref,args)
| _ -> None
-let constr_val_discr_st (idpred,cpred) t =
- let c, l = decomp t in
- match kind_of_term c with
+let constr_val_discr_st sigma (idpred,cpred) t =
+ let c, l = decomp sigma t in
+ match EConstr.kind sigma c with
| Const (c,u) -> if Cpred.mem c cpred then Everything else Label(GRLabel (ConstRef c),l)
| Ind (ind_sp,u) -> Label(GRLabel (IndRef ind_sp),l)
| Construct (cstr_sp,u) -> Label(GRLabel (ConstructRef cstr_sp),l)
@@ -105,11 +106,11 @@ let bounded_constr_pat_discr_st st (t,depth) =
| None -> None
| Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
-let bounded_constr_val_discr_st st (t,depth) =
+let bounded_constr_val_discr_st sigma st (t,depth) =
if Int.equal depth 0 then
Nothing
else
- match constr_val_discr_st st t with
+ match constr_val_discr_st sigma st t with
| Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l)
| Nothing -> Nothing
| Everything -> Everything
@@ -122,11 +123,11 @@ let bounded_constr_pat_discr (t,depth) =
| None -> None
| Some (c,l) -> Some(c,List.map (fun c -> (c,depth-1)) l)
-let bounded_constr_val_discr (t,depth) =
+let bounded_constr_val_discr sigma (t,depth) =
if Int.equal depth 0 then
Nothing
else
- match constr_val_discr t with
+ match constr_val_discr sigma t with
| Label (c,l) -> Label(c,List.map (fun c -> (c,depth-1)) l)
| Nothing -> Nothing
| Everything -> Everything
@@ -162,13 +163,13 @@ struct
(fun dn (c,v) ->
Dn.rmv dn (bounded_constr_pat_discr_st st) ((c,!dnet_depth),v))
- let lookup = function
+ let lookup sigma = function
| None ->
(fun dn t ->
- Dn.lookup dn bounded_constr_val_discr (t,!dnet_depth))
+ Dn.lookup dn (bounded_constr_val_discr sigma) (t,!dnet_depth))
| Some st ->
(fun dn t ->
- Dn.lookup dn (bounded_constr_val_discr_st st) (t,!dnet_depth))
+ Dn.lookup dn (bounded_constr_val_discr_st sigma st) (t,!dnet_depth))
let app f dn = Dn.app f dn
diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli
index 8ca5549b8..2a5e7c345 100644
--- a/tactics/btermdn.mli
+++ b/tactics/btermdn.mli
@@ -33,7 +33,7 @@ sig
val add : transparent_state option -> t -> (constr_pattern * Z.t) -> t
val rmv : transparent_state option -> t -> (constr_pattern * Z.t) -> t
- val lookup : transparent_state option -> t -> constr -> Z.t list
+ val lookup : Evd.evar_map -> transparent_state option -> t -> EConstr.constr -> Z.t list
val app : (Z.t -> unit) -> t -> unit
end
diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml
index edfe21d34..ea1966093 100644
--- a/tactics/class_tactics.ml
+++ b/tactics/class_tactics.ml
@@ -17,6 +17,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Reduction
open Proof_type
open Tacticals
@@ -189,8 +190,7 @@ let set_typeclasses_strategy = function
| Bfs -> set_typeclasses_iterative_deepening true
let pr_ev evs ev =
- Printer.pr_constr_env (Goal.V82.env evs ev) evs
- (Evarutil.nf_evar evs (Goal.V82.concl evs ev))
+ Printer.pr_econstr_env (Goal.V82.env evs ev) evs (Goal.V82.concl evs ev)
(** Typeclasses instance search tactic / eauto *)
@@ -236,13 +236,13 @@ let e_give_exact flags poly (c,clenv) gl =
let unify_e_resolve poly flags = { enter = begin fun gls (c,_,clenv) ->
let clenv', c = connect_hint_clenv poly c clenv gls in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in
+ let clenv' = clenv_unique_resolver ~flags clenv' gls in
Clenvtac.clenv_refine true ~with_classes:false clenv'
end }
let unify_resolve poly flags = { enter = begin fun gls (c,_,clenv) ->
let clenv', _ = connect_hint_clenv poly c clenv gls in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags clenv') gls in
+ let clenv' = clenv_unique_resolver ~flags clenv' gls in
Clenvtac.clenv_refine false ~with_classes:false clenv'
end }
@@ -264,7 +264,7 @@ let unify_resolve_refine poly flags gls ((c, t, ctx),n,clenv) =
sigma, c, t
in
let sigma', cl = Clenv.make_evar_clause env sigma ?len:n ty in
- let term = applistc term (List.map (fun x -> x.hole_evar) cl.cl_holes) in
+ let term = applist (term, List.map (fun x -> x.hole_evar) cl.cl_holes) in
let sigma' =
Evarconv.the_conv_x_leq env ~ts:flags.core_unify_flags.modulo_delta
cl.cl_concl concl sigma'
@@ -288,24 +288,24 @@ let clenv_of_prods poly nprods (c, clenv) gl =
let (c, _, _) = c in
if poly || Int.equal nprods 0 then Some (None, clenv)
else
- let ty = Retyping.get_type_of (Proofview.Goal.env gl)
- (Sigma.to_evar_map (Proofview.Goal.sigma gl)) c in
- let diff = nb_prod ty - nprods in
+ let sigma = Tacmach.New.project gl in
+ let ty = Retyping.get_type_of (Proofview.Goal.env gl) sigma c in
+ let diff = nb_prod sigma ty - nprods in
if Pervasives.(>=) diff 0 then
(* Was Some clenv... *)
Some (Some diff,
- Tacmach.New.of_old (fun gls -> mk_clenv_from_n gls (Some diff) (c,ty)) gl)
+ mk_clenv_from_n gl (Some diff) (c,ty))
else None
let with_prods nprods poly (c, clenv) f =
if get_typeclasses_limit_intros () then
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
try match clenv_of_prods poly nprods (c, clenv) gl with
| None -> Tacticals.New.tclZEROMSG (str"Not enough premisses")
| Some (diff, clenv') -> f.enter gl (c, diff, clenv')
with e when CErrors.noncritical e ->
Tacticals.New.tclZEROMSG (CErrors.print e) end }
- else Proofview.Goal.nf_enter
+ else Proofview.Goal.enter
{ enter = begin fun gl ->
if Int.equal nprods 0 then f.enter gl (c, None, clenv)
else Tacticals.New.tclZEROMSG (str"Not enough premisses") end }
@@ -351,21 +351,21 @@ let shelve_dependencies gls =
Feedback.msg_debug (str" shelving dependent subgoals: " ++ pr_gls sigma gls);
shelve_goals gls)
-let hintmap_of hdc secvars concl =
+let hintmap_of sigma hdc secvars concl =
match hdc with
| None -> fun db -> Hint_db.map_none secvars db
| Some hdc ->
fun db ->
if Hint_db.use_dn db then (* Using dnet *)
- Hint_db.map_eauto secvars hdc concl db
- else Hint_db.map_existential secvars hdc concl db
+ Hint_db.map_eauto sigma secvars hdc concl db
+ else Hint_db.map_existential sigma secvars hdc concl db
(** Hack to properly solve dependent evars that are typeclasses *)
let rec e_trivial_fail_db only_classes db_list local_db secvars =
let open Tacticals.New in
let open Tacmach.New in
let trivial_fail =
- Proofview.Goal.nf_enter { enter =
+ Proofview.Goal.enter { enter =
begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
@@ -376,7 +376,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars =
end }
in
let trivial_resolve =
- Proofview.Goal.nf_enter { enter =
+ Proofview.Goal.enter { enter =
begin fun gl ->
let tacs = e_trivial_resolve db_list local_db secvars only_classes
(project gl) (pf_concl gl) in
@@ -391,7 +391,7 @@ let rec e_trivial_fail_db only_classes db_list local_db secvars =
and e_my_find_search db_list local_db secvars hdc complete only_classes sigma concl =
let open Proofview.Notations in
- let prods, concl = decompose_prod_assum concl in
+ let prods, concl = EConstr.decompose_prod_assum sigma concl in
let nprods = List.length prods in
let freeze =
try
@@ -399,12 +399,12 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co
| Some (hd,_) when only_classes ->
let cl = Typeclasses.class_info hd in
if cl.cl_strict then
- Evd.evars_of_term concl
+ Evarutil.undefined_evars_of_term sigma concl
else Evar.Set.empty
| _ -> Evar.Set.empty
with e when CErrors.noncritical e -> Evar.Set.empty
in
- let hint_of_db = hintmap_of hdc secvars concl in
+ let hint_of_db = hintmap_of sigma hdc secvars concl in
let hintl =
List.map_append
(fun db ->
@@ -480,13 +480,13 @@ and e_my_find_search db_list local_db secvars hdc complete only_classes sigma co
in List.map tac_of_hint hintl
and e_trivial_resolve db_list local_db secvars only_classes sigma concl =
- let hd = try Some (decompose_app_bound concl) with Bound -> None in
+ let hd = try Some (decompose_app_bound sigma concl) with Bound -> None in
try
e_my_find_search db_list local_db secvars hd true only_classes sigma concl
with Not_found -> []
let e_possible_resolve db_list local_db secvars only_classes sigma concl =
- let hd = try Some (decompose_app_bound concl) with Bound -> None in
+ let hd = try Some (decompose_app_bound sigma concl) with Bound -> None in
try
e_my_find_search db_list local_db secvars hd false only_classes sigma concl
with Not_found -> []
@@ -511,13 +511,17 @@ let pr_depth l = prlist_with_sep (next_sep debug_seps) int (List.rev l)
let is_Prop env sigma concl =
let ty = Retyping.get_type_of env sigma concl in
- match kind_of_term ty with
- | Sort (Prop Null) -> true
+ match EConstr.kind sigma ty with
+ | Sort s ->
+ begin match ESorts.kind sigma s with
+ | Prop Null -> true
+ | _ -> false
+ end
| _ -> false
-let is_unique env concl =
+let is_unique env sigma concl =
try
- let (cl,u), args = dest_class_app env concl in
+ let (cl,u), args = dest_class_app env sigma concl in
cl.cl_unique
with e when CErrors.noncritical e -> false
@@ -560,14 +564,14 @@ let make_resolve_hyp env sigma st flags only_classes pri decl =
let id = NamedDecl.get_id decl in
let cty = Evarutil.nf_evar sigma (NamedDecl.get_type decl) in
let rec iscl env ty =
- let ctx, ar = decompose_prod_assum ty in
- match kind_of_term (fst (decompose_app ar)) with
+ let ctx, ar = decompose_prod_assum sigma ty in
+ match EConstr.kind sigma (fst (decompose_app sigma ar)) with
| Const (c,_) -> is_class (ConstRef c)
| Ind (i,_) -> is_class (IndRef i)
| _ ->
- let env' = Environ.push_rel_context ctx env in
- let ty' = whd_all env' ar in
- if not (Term.eq_constr ty' ar) then iscl env' ty'
+ let env' = push_rel_context ctx env in
+ let ty' = Reductionops.whd_all env' sigma ar in
+ if not (EConstr.eq_constr sigma ty' ar) then iscl env' ty'
else false
in
let is_class = iscl env cty in
@@ -587,7 +591,7 @@ let make_resolve_hyp env sigma st flags only_classes pri decl =
in
make_resolves env sigma ~name:(PathHints path)
(true,false,Flags.is_verbose()) info false
- (IsConstr (c,Univ.ContextSet.empty)))
+ (IsConstr (EConstr.of_constr c,Univ.ContextSet.empty)))
hints)
else []
in
@@ -605,7 +609,7 @@ let make_hints g st only_classes sign =
let consider =
try let t = hyp |> NamedDecl.get_id |> Global.lookup_named |> NamedDecl.get_type in
(* Section variable, reindex only if the type changed *)
- not (Term.eq_constr t (NamedDecl.get_type hyp))
+ not (EConstr.eq_constr (project g) (EConstr.of_constr t) (NamedDecl.get_type hyp))
with Not_found -> true
in
if consider then
@@ -660,7 +664,7 @@ module V85 = struct
then
cached_hints
else
- let hints = make_hints g st only_classes (Environ.named_context_of_val sign)
+ let hints = make_hints g st only_classes (EConstr.named_context_of_val sign)
in
cache := (only_classes, sign, hints); hints
@@ -677,7 +681,7 @@ module V85 = struct
let gls' =
List.map (fun g' ->
let env = Goal.V82.env s g' in
- let context = Environ.named_context_of_val (Goal.V82.hyps s g') in
+ let context = EConstr.named_context_of_val (Goal.V82.hyps s g') in
let hint = make_resolve_hyp env s (Hint_db.transparent_state info.hints)
(true,false,false) info.only_classes empty_hint_info (List.hd context) in
let ldb = Hint_db.add_list env s hint info.hints in
@@ -709,7 +713,7 @@ module V85 = struct
let needs_backtrack env evd oev concl =
if Option.is_empty oev || is_Prop env evd concl then
- occur_existential concl
+ occur_existential evd concl
else true
let hints_tac hints sk fk {it = gl,info; sigma = s} =
@@ -718,7 +722,7 @@ module V85 = struct
let tacgl = {it = gl; sigma = s;} in
let secvars = secvars_of_hyps (Environ.named_context_of_val (Goal.V82.hyps s gl)) in
let poss = e_possible_resolve hints info.hints secvars info.only_classes s concl in
- let unique = is_unique env concl in
+ let unique = is_unique env s concl in
let rec aux i foundone = function
| (tac, _, extern, name, pp) :: tl ->
let derivs = path_derivate info.auto_cut name in
@@ -783,10 +787,10 @@ module V85 = struct
let fk' =
(fun e ->
let do_backtrack =
- if unique then occur_existential concl
+ if unique then occur_existential tacgl.sigma concl
else if info.unique then true
else if List.is_empty gls' then
- needs_backtrack env s' info.is_evar concl
+ needs_backtrack env tacgl.sigma info.is_evar concl
else true
in
let e' = match foundone with None -> e
@@ -804,7 +808,7 @@ module V85 = struct
if foundone == None && !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.auto_depth ++ str": no match for " ++
- Printer.pr_constr_env (Goal.V82.env s gl) s concl ++
+ Printer.pr_econstr_env (Goal.V82.env s gl) s concl ++
spc () ++ str ", " ++ int (List.length poss) ++
str" possibilities");
match foundone with
@@ -987,13 +991,14 @@ module Search = struct
let sign = Goal.hyps g in
let (dir, onlyc, sign', cached_hints) = !autogoal_cache in
let cwd = Lib.cwd () in
+ let eq c1 c2 = EConstr.eq_constr (project g) c1 c2 in
if DirPath.equal cwd dir &&
(onlyc == only_classes) &&
- Context.Named.equal sign sign' &&
+ Context.Named.equal eq sign sign' &&
Hint_db.transparent_state cached_hints == st
then cached_hints
else
- let hints = make_hints {it = Goal.goal g; sigma = project g}
+ let hints = make_hints {it = Goal.goal (Proofview.Goal.assume g); sigma = project g}
st only_classes sign
in
autogoal_cache := (cwd, only_classes, sign, hints); hints
@@ -1025,7 +1030,7 @@ module Search = struct
NOT backtrack. *)
let needs_backtrack env evd unique concl =
if unique || is_Prop env evd concl then
- occur_existential concl
+ occur_existential evd concl
else true
let mark_unresolvables sigma goals =
@@ -1059,12 +1064,12 @@ module Search = struct
let concl = Goal.concl gl in
let sigma = Goal.sigma gl in
let s = Sigma.to_evar_map sigma in
- let unique = not info.search_dep || is_unique env concl in
+ let unique = not info.search_dep || is_unique env s concl in
let backtrack = needs_backtrack env s unique concl in
if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.search_depth ++ str": looking for " ++
- Printer.pr_constr_env (Goal.env gl) s concl ++
+ Printer.pr_econstr_env (Goal.env gl) s concl ++
(if backtrack then str" with backtracking"
else str" without backtracking"));
let secvars = compute_secvars gl in
@@ -1086,28 +1091,29 @@ module Search = struct
pr_depth (!idx :: info.search_depth) ++ str": " ++
Lazy.force pp ++
(if !foundone != true then
- str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl)
+ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl))
else mt ())
in
Feedback.msg_debug (msg ++ str " failed with " ++ CErrors.iprint ie)
else ()
in
- let tac_of gls i j = Goal.nf_enter { enter = fun gl' ->
+ let tac_of gls i j = Goal.enter { enter = fun gl' ->
let sigma' = Goal.sigma gl' in
let s' = Sigma.to_evar_map sigma' in
let _concl = Goal.concl gl' in
if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth (succ j :: i :: info.search_depth) ++ str" : " ++
- pr_ev s' (Proofview.Goal.goal gl'));
+ pr_ev s' (Proofview.Goal.goal (Proofview.Goal.assume gl')));
+ let eq c1 c2 = EConstr.eq_constr s' c1 c2 in
let hints' =
- if b && not (Context.Named.equal (Goal.hyps gl') (Goal.hyps gl))
+ if b && not (Context.Named.equal eq (Goal.hyps gl') (Goal.hyps gl))
then
let st = Hint_db.transparent_state info.search_hints in
make_autogoal_hints info.search_only_classes ~st gl'
else info.search_hints
in
- let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal gl') gls in
+ let dep' = info.search_dep || Proofview.unifiable s' (Goal.goal (Proofview.Goal.assume gl')) gls in
let info' =
{ search_depth = succ j :: i :: info.search_depth;
last_tac = pp;
@@ -1124,7 +1130,7 @@ module Search = struct
(if !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth (i :: info.search_depth) ++ str": " ++ Lazy.force pp
- ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal gl)
+ ++ str" on" ++ spc () ++ pr_ev s (Proofview.Goal.goal (Proofview.Goal.assume gl))
++ str", " ++ int j ++ str" subgoal(s)" ++
(Option.cata (fun k -> str " in addition to the first " ++ int k)
(mt()) k)));
@@ -1138,7 +1144,7 @@ module Search = struct
try
let evi = Evd.find_undefined sigma ev in
if info.search_only_classes then
- Some (ev, not (is_class_type sigma (Evd.evar_concl evi)))
+ Some (ev, not (is_class_evar sigma evi))
else Some (ev, true)
with Not_found -> None
in
@@ -1194,7 +1200,7 @@ module Search = struct
if !foundone == false && !typeclasses_debug > 0 then
Feedback.msg_debug
(pr_depth info.search_depth ++ str": no match for " ++
- Printer.pr_constr_env (Goal.env gl) s concl ++
+ Printer.pr_econstr_env (Goal.env gl) s concl ++
spc () ++ str ", " ++ int (List.length poss) ++
str" possibilities");
match e with
@@ -1205,7 +1211,7 @@ module Search = struct
else tclONCE (aux (NotApplicableEx,Exninfo.null) poss)
let hints_tac hints info kont : unit Proofview.tactic =
- Proofview.Goal.nf_enter
+ Proofview.Goal.enter
{ enter = fun gl -> hints_tac_gl hints info kont gl }
let intro_tac info kont gl =
@@ -1226,7 +1232,7 @@ module Search = struct
let intro info kont =
Proofview.tclBIND Tactics.intro
- (fun _ -> Proofview.Goal.nf_enter { enter = fun gl -> intro_tac info kont gl })
+ (fun _ -> Proofview.Goal.enter { enter = fun gl -> intro_tac info kont gl })
let rec search_tac hints limit depth =
let kont info =
@@ -1252,14 +1258,14 @@ module Search = struct
if false (* In 8.6, still allow non-class goals only_classes && not (is_class_type sigma (Goal.concl gl)) *) then
Tacticals.New.tclZEROMSG (str"Not a subgoal for a class")
else
- let dep = dep || Proofview.unifiable sigma (Goal.goal gl) gls in
+ let dep = dep || Proofview.unifiable sigma (Goal.goal (Proofview.Goal.assume gl)) gls in
let info = make_autogoal ?st only_classes dep (cut_of_hints hints) i gl in
search_tac hints depth 1 info
let search_tac ?(st=full_transparent_state) only_classes dep hints depth =
let open Proofview in
let tac sigma gls i =
- Goal.nf_enter
+ Goal.enter
{ enter = fun gl ->
search_tac_gl ~st only_classes dep hints depth (succ i) sigma gls gl }
in
@@ -1410,8 +1416,8 @@ module Intpart = Unionfind.Make(Evar.Set)(Evar.Map)
let deps_of_constraints cstrs evm p =
List.iter (fun (_, _, x, y) ->
- let evx = Evarutil.undefined_evars_of_term evm x in
- let evy = Evarutil.undefined_evars_of_term evm y in
+ let evx = Evarutil.undefined_evars_of_term evm (EConstr.of_constr x) in
+ let evy = Evarutil.undefined_evars_of_term evm (EConstr.of_constr y) in
Intpart.union_set (Evar.Set.union evx evy) p)
cstrs
@@ -1455,7 +1461,7 @@ let error_unresolvable env comp evd =
| Some s -> Evar.Set.mem ev s
in
let fold ev evi (found, accu) =
- let ev_class = class_of_constr evi.evar_concl in
+ let ev_class = class_of_constr evd (EConstr.of_constr evi.evar_concl) in
if not (Option.is_empty ev_class) && is_part ev then
(* focus on one instance if only one was searched for *)
if not found then (true, Some ev)
@@ -1560,9 +1566,10 @@ let _ =
Hook.set Typeclasses.solve_all_instances_hook solve_inst
let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique =
- let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env gl in
+ let nc, gl, subst, _, _ = Evarutil.push_rel_context_to_named_context env sigma gl in
let (gl,t,sigma) =
Goal.V82.mk_goal sigma nc gl Store.empty in
+ let (ev, _) = destEvar sigma t in
let gls = { it = gl ; sigma = sigma; } in
let hints = searchtable_map typeclasses_db in
let st = Hint_db.transparent_state hints in
@@ -1577,9 +1584,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl unique =
with Refiner.FailError _ -> raise Not_found
in
let evd = sig_sig gls' in
- let t' = let (ev, inst) = destEvar t in
- mkEvar (ev, Array.of_list subst)
- in
+ let t' = mkEvar (ev, Array.of_list subst) in
let term = Evarutil.nf_evar evd t' in
evd, term
@@ -1590,21 +1595,22 @@ let _ =
(** Take the head of the arity of a constr.
Used in the partial application tactic. *)
-let rec head_of_constr t =
- let t = strip_outer_cast(collapse_appl t) in
- match kind_of_term t with
- | Prod (_,_,c2) -> head_of_constr c2
- | LetIn (_,_,_,c2) -> head_of_constr c2
- | App (f,args) -> head_of_constr f
+let rec head_of_constr sigma t =
+ let t = strip_outer_cast sigma (collapse_appl sigma t) in
+ match EConstr.kind sigma t with
+ | Prod (_,_,c2) -> head_of_constr sigma c2
+ | LetIn (_,_,_,c2) -> head_of_constr sigma c2
+ | App (f,args) -> head_of_constr sigma f
| _ -> t
let head_of_constr h c =
- let c = head_of_constr c in
+ Proofview.tclEVARMAP >>= fun sigma ->
+ let c = head_of_constr sigma c in
letin_tac None (Name h) c None Locusops.allHyps
let not_evar c =
Proofview.tclEVARMAP >>= fun sigma ->
- match Evarutil.kind_of_term_upto sigma c with
+ match EConstr.kind sigma c with
| Evar _ -> Tacticals.New.tclFAIL 0 (str"Evar")
| _ -> Proofview.tclUNIT ()
@@ -1612,17 +1618,15 @@ let is_ground c gl =
if Evarutil.is_ground_term (project gl) c then tclIDTAC gl
else tclFAIL 0 (str"Not ground") gl
-let autoapply c i gl =
+let autoapply c i =
let open Proofview.Notations in
+ Proofview.Goal.enter { enter = begin fun gl ->
let flags = auto_unif_flags Evar.Set.empty
(Hints.Hint_db.transparent_state (Hints.searchtable_map i)) in
- let cty = pf_unsafe_type_of gl c in
+ let cty = Tacmach.New.pf_unsafe_type_of gl c in
let ce = mk_clenv_from gl (c,cty) in
- let enter gl =
(unify_e_resolve false flags).enter gl
((c,cty,Univ.ContextSet.empty),0,ce) <*>
Proofview.tclEVARMAP >>= (fun sigma ->
let sigma = Typeclasses.mark_unresolvables ~filter:Typeclasses.all_goals sigma in
- Proofview.Unsafe.tclEVARS sigma)
- in
- Proofview.V82.of_tactic (Proofview.Goal.nf_enter { enter }) gl
+ Proofview.Unsafe.tclEVARS sigma) end }
diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli
index 76760db02..a38be5972 100644
--- a/tactics/class_tactics.mli
+++ b/tactics/class_tactics.mli
@@ -10,6 +10,7 @@
open Names
open Constr
+open EConstr
open Tacmach
val catchable : exn -> bool
@@ -28,13 +29,13 @@ val typeclasses_eauto : ?only_classes:bool -> ?st:transparent_state -> ?strategy
depth:(Int.t option) ->
Hints.hint_db_name list -> unit Proofview.tactic
-val head_of_constr : Id.t -> Term.constr -> unit Proofview.tactic
+val head_of_constr : Id.t -> constr -> unit Proofview.tactic
val not_evar : constr -> unit Proofview.tactic
val is_ground : constr -> tactic
-val autoapply : constr -> Hints.hint_db_name -> tactic
+val autoapply : constr -> Hints.hint_db_name -> unit Proofview.tactic
module Search : sig
val eauto_tac :
diff --git a/tactics/contradiction.ml b/tactics/contradiction.ml
index 6b29f574c..63f923dfd 100644
--- a/tactics/contradiction.ml
+++ b/tactics/contradiction.ml
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Hipattern
open Tactics
open Coqlib
@@ -19,6 +20,7 @@ module NamedDecl = Context.Named.Declaration
(* Absurd *)
let mk_absurd_proof t =
+ let build_coq_not () = EConstr.of_constr (build_coq_not ()) in
let id = Namegen.default_dependent_ident in
mkLambda (Names.Name id,mkApp(build_coq_not (),[|t|]),
mkLambda (Names.Name id,t,mkApp (mkRel 2,[|mkRel 1|])))
@@ -33,7 +35,7 @@ let absurd c =
let t = j.Environ.utj_val in
let tac =
Tacticals.New.tclTHENLIST [
- elim_type (build_coq_False ());
+ elim_type (EConstr.of_constr (build_coq_False ()));
Simple.apply (mk_absurd_proof t)
] in
Sigma.Unsafe.of_pair (tac, sigma)
@@ -66,18 +68,18 @@ let contradiction_context =
let id = NamedDecl.get_id d in
let typ = nf_evar sigma (NamedDecl.get_type d) in
let typ = whd_all env sigma typ in
- if is_empty_type typ then
+ if is_empty_type sigma typ then
simplest_elim (mkVar id)
- else match kind_of_term typ with
- | Prod (na,t,u) when is_empty_type u ->
+ else match EConstr.kind sigma typ with
+ | Prod (na,t,u) when is_empty_type sigma u ->
let is_unit_or_eq =
- if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type t
+ if use_negated_unit_or_eq_type () then match_with_unit_or_eq_type sigma t
else None in
Tacticals.New.tclORELSE
(match is_unit_or_eq with
| Some _ ->
- let hd,args = decompose_app t in
- let (ind,_ as indu) = destInd hd in
+ let hd,args = decompose_app sigma t in
+ let (ind,_ as indu) = destInd sigma hd in
let nparams = Inductiveops.inductive_nparams_env env ind in
let params = Util.List.firstn nparams args in
let p = applist ((mkConstructUi (indu,1)), params) in
@@ -102,20 +104,19 @@ let contradiction_context =
end }
let is_negation_of env sigma typ t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (na,t,u) ->
- let u = nf_evar sigma u in
- is_empty_type u && is_conv_leq env sigma typ t
+ is_empty_type sigma u && is_conv_leq env sigma typ t
| _ -> false
let contradiction_term (c,lbind as cl) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
let type_of = Tacmach.New.pf_unsafe_type_of gl in
let typ = type_of c in
let _, ccl = splay_prod env sigma typ in
- if is_empty_type ccl then
+ if is_empty_type sigma ccl then
Tacticals.New.tclTHEN
(elim false None cl None)
(Tacticals.New.tclTRY assumption)
@@ -123,7 +124,7 @@ let contradiction_term (c,lbind as cl) =
Proofview.tclORELSE
begin
if lbind = NoBindings then
- filter_hyp (is_negation_of env sigma typ)
+ filter_hyp (fun c -> is_negation_of env sigma typ c)
(fun id -> simplest_elim (mkApp (mkVar id,[|c|])))
else
Proofview.tclZERO Not_found
diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli
index b876aee90..510b135b0 100644
--- a/tactics/contradiction.mli
+++ b/tactics/contradiction.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Misctypes
val absurd : constr -> unit Proofview.tactic
diff --git a/tactics/eauto.ml b/tactics/eauto.ml
index 885183174..8d1e0e507 100644
--- a/tactics/eauto.ml
+++ b/tactics/eauto.ml
@@ -13,6 +13,7 @@ open Names
open Nameops
open Term
open Termops
+open EConstr
open Proof_type
open Tacticals
open Tacmach
@@ -32,7 +33,8 @@ let e_give_exact ?(flags=eauto_unif_flags) c =
Proofview.Goal.enter { enter = begin fun gl ->
let t1 = Tacmach.New.pf_unsafe_type_of gl c in
let t2 = Tacmach.New.pf_concl (Proofview.Goal.assume gl) in
- if occur_existential t1 || occur_existential t2 then
+ let sigma = Tacmach.New.project gl in
+ if occur_existential sigma t1 || occur_existential sigma t2 then
Tacticals.New.tclTHEN (Clenvtac.unify ~flags t1) (exact_no_check c)
else exact_check c
end }
@@ -86,7 +88,7 @@ let rec prolog l n gl =
let out_term = function
| IsConstr (c, _) -> c
- | IsGlobRef gr -> fst (Universes.fresh_global_instance (Global.env ()) gr)
+ | IsGlobRef gr -> EConstr.of_constr (fst (Universes.fresh_global_instance (Global.env ()) gr))
let prolog_tac l n =
Proofview.V82.tactic begin fun gl ->
@@ -110,22 +112,21 @@ open Auto
let priority l = List.map snd (List.filter (fun (pr,_) -> Int.equal pr 0) l)
let unify_e_resolve poly flags (c,clenv) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let clenv', c = connect_hint_clenv poly c clenv gl in
- Proofview.V82.tactic
- (fun gls ->
- let clenv' = clenv_unique_resolver ~flags clenv' gls in
- tclTHEN (Refiner.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd))
- (Proofview.V82.of_tactic (Tactics.Simple.eapply c)) gls)
+ let clenv' = clenv_unique_resolver ~flags clenv' gl in
+ Proofview.tclTHEN
+ (Proofview.Unsafe.tclEVARUNIVCONTEXT (Evd.evar_universe_context clenv'.evd))
+ (Tactics.Simple.eapply c)
end }
-let hintmap_of secvars hdc concl =
+let hintmap_of sigma secvars hdc concl =
match hdc with
| None -> fun db -> Hint_db.map_none ~secvars db
| Some hdc ->
- if occur_existential concl then
- (fun db -> Hint_db.map_existential ~secvars hdc concl db)
- else (fun db -> Hint_db.map_auto ~secvars hdc concl db)
+ if occur_existential sigma concl then
+ (fun db -> Hint_db.map_existential sigma ~secvars hdc concl db)
+ else (fun db -> Hint_db.map_auto sigma ~secvars hdc concl db)
(* FIXME: should be (Hint_db.map_eauto hdc concl db) *)
let e_exact poly flags (c,clenv) =
@@ -137,7 +138,7 @@ let e_exact poly flags (c,clenv) =
end }
let rec e_trivial_fail_db db_list local_db =
- let next = Proofview.Goal.nf_enter { enter = begin fun gl ->
+ let next = Proofview.Goal.enter { enter = begin fun gl ->
let d = Tacmach.New.pf_last_hyp gl in
let hintl = make_resolve_hyp (Tacmach.New.pf_env gl) (Tacmach.New.project gl) d in
e_trivial_fail_db db_list (Hint_db.add_list (Tacmach.New.pf_env gl) (Tacmach.New.project gl) hintl local_db)
@@ -147,13 +148,13 @@ let rec e_trivial_fail_db db_list local_db =
let tacl =
registered_e_assumption ::
(Tacticals.New.tclTHEN Tactics.intro next) ::
- (List.map fst (e_trivial_resolve db_list local_db secvars (Tacmach.New.pf_nf_concl gl)))
+ (List.map fst (e_trivial_resolve (Tacmach.New.project gl) db_list local_db secvars (Tacmach.New.pf_concl gl)))
in
Tacticals.New.tclFIRST (List.map Tacticals.New.tclCOMPLETE tacl)
end }
-and e_my_find_search db_list local_db secvars hdc concl =
- let hint_of_db = hintmap_of secvars hdc concl in
+and e_my_find_search sigma db_list local_db secvars hdc concl =
+ let hint_of_db = hintmap_of sigma secvars hdc concl in
let hintl =
List.map_append (fun db ->
let flags = auto_flags_of_state (Hint_db.transparent_state db) in
@@ -181,15 +182,15 @@ and e_my_find_search db_list local_db secvars hdc concl =
in
List.map tac_of_hint hintl
-and e_trivial_resolve db_list local_db secvars gl =
- let hd = try Some (decompose_app_bound gl) with Bound -> None in
- try priority (e_my_find_search db_list local_db secvars hd gl)
+and e_trivial_resolve sigma db_list local_db secvars gl =
+ let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in
+ try priority (e_my_find_search sigma db_list local_db secvars hd gl)
with Not_found -> []
-let e_possible_resolve db_list local_db secvars gl =
- let hd = try Some (decompose_app_bound gl) with Bound -> None in
+let e_possible_resolve sigma db_list local_db secvars gl =
+ let hd = try Some (decompose_app_bound sigma gl) with Bound -> None in
try List.map (fun (b, (tac, pp)) -> (tac, b, pp))
- (e_my_find_search db_list local_db secvars hd gl)
+ (e_my_find_search sigma db_list local_db secvars hd gl)
with Not_found -> []
let find_first_goal gls =
@@ -287,9 +288,9 @@ module SearchProblem = struct
in
let rec_tacs =
let l =
- let concl = Reductionops.nf_evar (project g)(pf_concl g) in
+ let concl = Reductionops.nf_evar (project g) (pf_concl g) in
filter_tactics s.tacres
- (e_possible_resolve s.dblist (List.hd s.localdb) secvars concl)
+ (e_possible_resolve (project g) s.dblist (List.hd s.localdb) secvars concl)
in
List.map
(fun (lgls, cost, pp) ->
@@ -464,18 +465,19 @@ let autounfold_tac db cls =
in
autounfold dbs cls
-let unfold_head env (ids, csts) c =
+let unfold_head env sigma (ids, csts) c =
let rec aux c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id when Id.Set.mem id ids ->
(match Environ.named_body id env with
- | Some b -> true, b
+ | Some b -> true, EConstr.of_constr b
| None -> false, c)
- | Const (cst,u as c) when Cset.mem cst csts ->
- true, Environ.constant_value_in env c
+ | Const (cst, u) when Cset.mem cst csts ->
+ let u = EInstance.kind sigma u in
+ true, EConstr.of_constr (Environ.constant_value_in env (cst, u))
| App (f, args) ->
(match aux f with
- | true, f' -> true, Reductionops.whd_betaiota Evd.empty (mkApp (f', args))
+ | true, f' -> true, Reductionops.whd_betaiota sigma (mkApp (f', args))
| false, _ ->
let done_, args' =
Array.fold_left_i (fun i (done_, acc) arg ->
@@ -489,7 +491,7 @@ let unfold_head env (ids, csts) c =
else false, c)
| _ ->
let done_ = ref false in
- let c' = map_constr (fun c ->
+ let c' = EConstr.map sigma (fun c ->
if !done_ then c else
let x, c' = aux c in
done_ := x; c') c
@@ -497,8 +499,9 @@ let unfold_head env (ids, csts) c =
in aux c
let autounfold_one db cl =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
let st =
List.fold_left (fun (i,c) dbname ->
@@ -508,7 +511,7 @@ let autounfold_one db cl =
let (ids, csts) = Hint_db.unfolds db in
(Id.Set.union ids i, Cset.union csts c)) (Id.Set.empty, Cset.empty) db
in
- let did, c' = unfold_head env st
+ let did, c' = unfold_head env sigma st
(match cl with Some (id, _) -> Tacmach.New.pf_get_hyp_typ id gl | None -> concl)
in
if did then
diff --git a/tactics/eauto.mli b/tactics/eauto.mli
index 1f69e4ab3..e2006ac1e 100644
--- a/tactics/eauto.mli
+++ b/tactics/eauto.mli
@@ -7,6 +7,7 @@
(************************************************************************)
open Term
+open EConstr
open Proof_type
open Hints
open Tactypes
diff --git a/tactics/elim.ml b/tactics/elim.ml
index 3f0c01a29..e37ec6bce 100644
--- a/tactics/elim.ml
+++ b/tactics/elim.ml
@@ -10,6 +10,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Inductiveops
open Hipattern
open Tacmach.New
@@ -79,11 +80,12 @@ let up_to_delta = ref false (* true *)
let general_decompose recognizer c =
Proofview.Goal.enter { enter = begin fun gl ->
let type_of = pf_unsafe_type_of gl in
+ let sigma = project gl in
let typc = type_of c in
tclTHENS (cut typc)
[ tclTHEN (intro_using tmphyp_name)
(onLastHypId
- (ifOnHyp recognizer (general_decompose_aux recognizer)
+ (ifOnHyp (recognizer sigma) (general_decompose_aux (recognizer sigma))
(fun id -> clear [id])));
exact_no_check c ]
end }
@@ -95,24 +97,24 @@ let head_in indl t gl =
let ity,_ =
if !up_to_delta
then find_mrectype env sigma t
- else extract_mrectype t
+ else extract_mrectype sigma t
in List.exists (fun i -> eq_ind (fst i) (fst ity)) indl
with Not_found -> false
let decompose_these c l =
Proofview.Goal.enter { enter = begin fun gl ->
let indl = List.map (fun x -> x, Univ.Instance.empty) l in
- general_decompose (fun (_,t) -> head_in indl t gl) c
+ general_decompose (fun sigma (_,t) -> head_in indl t gl) c
end }
let decompose_and c =
general_decompose
- (fun (_,t) -> is_record t)
+ (fun sigma (_,t) -> is_record sigma t)
c
let decompose_or c =
general_decompose
- (fun (_,t) -> is_disjunction t)
+ (fun sigma (_,t) -> is_disjunction sigma t)
c
let h_decompose l c = decompose_these c l
@@ -131,9 +133,9 @@ let induction_trailer abs_i abs_j bargs =
(tclDO (abs_j - abs_i) intro)
(onLastHypId
(fun id ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let idty = pf_unsafe_type_of gl (mkVar id) in
- let fvty = global_vars (pf_env gl) idty in
+ let fvty = global_vars (pf_env gl) (project gl) idty in
let possible_bring_hyps =
(List.tl (nLastDecls gl (abs_j - abs_i))) @ bargs.Tacticals.assums
in
@@ -153,7 +155,7 @@ let induction_trailer abs_i abs_j bargs =
))
let double_ind h1 h2 =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let abs_i = depth_of_quantified_hypothesis true h1 gl in
let abs_j = depth_of_quantified_hypothesis true h2 gl in
let abs =
diff --git a/tactics/elim.mli b/tactics/elim.mli
index 29c441463..dc1af79ba 100644
--- a/tactics/elim.mli
+++ b/tactics/elim.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Tacticals
open Misctypes
open Tactypes
diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml
index 1a67bedc2..bac3980d2 100644
--- a/tactics/eqdecide.ml
+++ b/tactics/eqdecide.ml
@@ -16,6 +16,7 @@ open Util
open Names
open Namegen
open Term
+open EConstr
open Declarations
open Tactics
open Tacticals.New
@@ -25,6 +26,7 @@ open Misctypes
open Tactypes
open Hipattern
open Pretyping
+open Proofview.Notations
open Tacmach.New
open Coqlib
@@ -50,7 +52,9 @@ open Coqlib
Eduardo Gimenez (30/3/98).
*)
-let clear_last = (onLastHyp (fun c -> (clear [destVar c])))
+let clear_last =
+ Proofview.tclEVARMAP >>= fun sigma ->
+ (onLastHyp (fun c -> (clear [destVar sigma c])))
let choose_eq eqonleft =
if eqonleft then
@@ -73,7 +77,7 @@ let mkBranches c1 c2 =
intros]
let discrHyp id =
- let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in
+ let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in
let tac c = Equality.discr_tac false (Some (None, ElimOnConstr c)) in
Tacticals.New.tclDELAYEDWITHHOLES false c tac
@@ -85,7 +89,9 @@ let solveNoteqBranch side =
(* Constructs the type {c1=c2}+{~c1=c2} *)
let make_eq () =
-(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+(*FIXME*) EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ()))
+let build_coq_not () = EConstr.of_constr (build_coq_not ())
+let build_coq_sumbool () = EConstr.of_constr (build_coq_sumbool ())
let mkDecideEqGoal eqonleft op rectype c1 c2 =
let equality = mkApp(make_eq(), [|rectype; c1; c2|]) in
@@ -121,7 +127,7 @@ let eqCase tac =
tclTHEN intro (onLastHypId tac)
let injHyp id =
- let c = { delayed = fun env sigma -> Sigma.here (Term.mkVar id, NoBindings) sigma } in
+ let c = { delayed = fun env sigma -> Sigma.here (mkVar id, NoBindings) sigma } in
let tac c = Equality.injClause None false (Some (None, ElimOnConstr c)) in
Tacticals.New.tclDELAYEDWITHHOLES false c tac
@@ -141,8 +147,8 @@ open Proofview.Notations
(* spiwack: a small wrapper around [Hipattern]. *)
-let match_eqdec c =
- try Proofview.tclUNIT (match_eqdec c)
+let match_eqdec sigma c =
+ try Proofview.tclUNIT (match_eqdec sigma c)
with PatternMatchingFailure -> Proofview.tclZERO PatternMatchingFailure
(* /spiwack *)
@@ -170,11 +176,12 @@ let solveEqBranch rectype =
Proofview.tclORELSE
begin
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
- match_eqdec concl >>= fun (eqonleft,op,lhs,rhs,_) ->
+ let concl = pf_concl gl in
+ let sigma = project gl in
+ match_eqdec sigma concl >>= fun (eqonleft,op,lhs,rhs,_) ->
let (mib,mip) = Global.lookup_inductive rectype in
let nparams = mib.mind_nparams in
- let getargs l = List.skipn nparams (snd (decompose_app l)) in
+ let getargs l = List.skipn nparams (snd (decompose_app sigma l)) in
let rargs = getargs rhs
and largs = getargs lhs in
solveArg [] eqonleft op largs rargs
@@ -187,7 +194,7 @@ let solveEqBranch rectype =
(* The tactic Decide Equality *)
-let hd_app c = match kind_of_term c with
+let hd_app sigma c = match EConstr.kind sigma c with
| App (h,_) -> h
| _ -> c
@@ -195,10 +202,11 @@ let decideGralEquality =
Proofview.tclORELSE
begin
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
- match_eqdec concl >>= fun (eqonleft,_,c1,c2,typ) ->
- let headtyp = hd_app (pf_compute gl typ) in
- begin match kind_of_term headtyp with
+ let concl = pf_concl gl in
+ let sigma = project gl in
+ match_eqdec sigma concl >>= fun (eqonleft,_,c1,c2,typ) ->
+ let headtyp = hd_app sigma (pf_compute gl typ) in
+ begin match EConstr.kind sigma headtyp with
| Ind (mi,_) -> Proofview.tclUNIT mi
| _ -> tclZEROMSG (Pp.str"This decision procedure only works for inductive objects.")
end >>= fun rectype ->
diff --git a/tactics/eqdecide.mli b/tactics/eqdecide.mli
index cb48a5bcc..dca1780b7 100644
--- a/tactics/eqdecide.mli
+++ b/tactics/eqdecide.mli
@@ -14,4 +14,4 @@
val decideEqualityGoal : unit Proofview.tactic
-val compare : Constr.t -> Constr.t -> unit Proofview.tactic
+val compare : EConstr.t -> EConstr.t -> unit Proofview.tactic
diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml
index c94dcfa9d..b08456f2f 100644
--- a/tactics/eqschemes.ml
+++ b/tactics/eqschemes.ml
@@ -73,13 +73,27 @@ let build_dependent_inductive ind (mib,mip) =
let realargs,_ = List.chop mip.mind_nrealdecls mip.mind_arity_ctxt in
applist
(mkIndU ind,
- Context.Rel.to_extended_list mip.mind_nrealdecls mib.mind_params_ctxt
- @ Context.Rel.to_extended_list 0 realargs)
+ Context.Rel.to_extended_list mkRel mip.mind_nrealdecls mib.mind_params_ctxt
+ @ Context.Rel.to_extended_list mkRel 0 realargs)
+
+let named_hd env t na = named_hd env Evd.empty (EConstr.of_constr t) na
+let name_assumption env = function
+| LocalAssum (na,t) -> LocalAssum (named_hd env t na, t)
+| LocalDef (na,c,t) -> LocalDef (named_hd env c na, c, t)
+
+let name_context env hyps =
+ snd
+ (List.fold_left
+ (fun (env,hyps) d ->
+ let d' = name_assumption env d in (push_rel d' env, d' :: hyps))
+ (env,[]) (List.rev hyps))
let my_it_mkLambda_or_LetIn s c = it_mkLambda_or_LetIn c s
-let my_it_mkProd_or_LetIn s c = it_mkProd_or_LetIn c s
+let my_it_mkProd_or_LetIn s c = Term.it_mkProd_or_LetIn c s
let my_it_mkLambda_or_LetIn_name s c =
- it_mkLambda_or_LetIn_name (Global.env()) c s
+ let env = Global.env () in
+ let mkLambda_or_LetIn_name d b = mkLambda_or_LetIn (name_assumption env d) b in
+ List.fold_left (fun c d -> mkLambda_or_LetIn_name d c) c s
let get_coq_eq ctx =
try
@@ -120,7 +134,7 @@ let get_sym_eq_data env (ind,u) =
let paramsctxt = Vars.subst_instance_context u mib.mind_params_ctxt in
let paramsctxt1,_ =
List.chop (mib.mind_nparams-mip.mind_nrealargs) paramsctxt in
- if not (List.equal eq_constr params2 constrargs) then
+ if not (List.equal Term.eq_constr params2 constrargs) then
error "Constructors arguments must repeat the parameters.";
(* nrealargs_ctxt and nrealargs are the same here *)
(specif,mip.mind_nrealargs,realsign,paramsctxt,paramsctxt1)
@@ -172,7 +186,7 @@ let build_sym_scheme env ind =
let (mib,mip as specif),nrealargs,realsign,paramsctxt,paramsctxt1 =
get_sym_eq_data env indu in
let cstr n =
- mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in
+ mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let applied_ind = build_dependent_inductive indu specif in
let realsign_ind =
@@ -185,7 +199,7 @@ let build_sym_scheme env ind =
my_it_mkLambda_or_LetIn_name
(lift_rel_context (nrealargs+1) realsign_ind)
(mkApp (mkIndU indu,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect 1 nrealargs;
rel_vect (2*nrealargs+2) nrealargs])),
mkRel 1 (* varH *),
@@ -226,13 +240,13 @@ let build_sym_involutive_scheme env ind =
get_sym_eq_data env indu in
let eq,eqrefl,ctx = get_coq_eq ctx in
let sym, ctx, eff = const_of_scheme sym_scheme_kind env ind ctx in
- let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect n paramsctxt) in
+ let cstr n = mkApp (mkConstructUi (indu,1),Context.Rel.to_extended_vect mkRel n paramsctxt) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_C =
mkApp
(mkIndU indu, Array.append
- (Context.Rel.to_extended_vect (nrealargs+1) mib.mind_params_ctxt)
+ (Context.Rel.to_extended_vect mkRel (nrealargs+1) mib.mind_params_ctxt)
(rel_vect (nrealargs+1) nrealargs)) in
let realsign_ind =
name_context env ((LocalAssum (Name varH,applied_ind))::realsign) in
@@ -246,15 +260,15 @@ let build_sym_involutive_scheme env ind =
(mkApp (eq,[|
mkApp
(mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect (2*nrealargs+2) nrealargs;
rel_vect 1 nrealargs]);
mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect 1 nrealargs;
rel_vect (2*nrealargs+2) nrealargs;
[|mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+2) paramsctxt1;
rel_vect (2*nrealargs+2) nrealargs;
rel_vect 1 nrealargs;
[|mkRel 1|]])|]]);
@@ -337,7 +351,7 @@ let build_l2r_rew_scheme dep env ind kind =
let eq,eqrefl,ctx = get_coq_eq ctx in
let cstr n p =
mkApp (mkConstructUi(indu,1),
- Array.concat [Context.Rel.to_extended_vect n paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -345,12 +359,12 @@ let build_l2r_rew_scheme dep env ind kind =
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_P =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect nrealargs nrealargs]) in
let applied_ind_G =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+3) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+3) paramsctxt1;
rel_vect (nrealargs+3) nrealargs;
rel_vect 0 nrealargs]) in
let realsign_P = lift_rel_context nrealargs realsign in
@@ -361,10 +375,10 @@ let build_l2r_rew_scheme dep env ind kind =
lift_rel_context (nrealargs+3) realsign) in
let applied_sym_C n =
mkApp(sym,
- Array.append (Context.Rel.to_extended_vect n mip.mind_arity_ctxt) [|mkVar varH|]) in
+ Array.append (Context.Rel.to_extended_vect mkRel n mip.mind_arity_ctxt) [|mkVar varH|]) in
let applied_sym_G =
mkApp(sym,
- Array.concat [Context.Rel.to_extended_vect (nrealargs*3+4) paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel (nrealargs*3+4) paramsctxt1;
rel_vect (nrealargs+4) nrealargs;
rel_vect 1 nrealargs;
[|mkRel 1|]]) in
@@ -374,7 +388,7 @@ let build_l2r_rew_scheme dep env ind kind =
let ci = make_case_info (Global.env()) ind RegularStyle in
let cieq = make_case_info (Global.env()) (fst (destInd eq)) RegularStyle in
let applied_PC =
- mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect 1 realsign)
+ mkApp (mkVar varP,Array.append (Context.Rel.to_extended_vect mkRel 1 realsign)
(if dep then [|cstr (2*nrealargs+1) 1|] else [||])) in
let applied_PG =
mkApp (mkVar varP,Array.append (rel_vect 1 nrealargs)
@@ -384,11 +398,11 @@ let build_l2r_rew_scheme dep env ind kind =
(if dep then [|mkRel 2|] else [||])) in
let applied_sym_sym =
mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1;
rel_vect 4 nrealargs;
rel_vect (nrealargs+4) nrealargs;
[|mkApp (sym,Array.concat
- [Context.Rel.to_extended_vect (2*nrealargs+4) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (2*nrealargs+4) paramsctxt1;
rel_vect (nrealargs+4) nrealargs;
rel_vect 4 nrealargs;
[|mkRel 2|]])|]]) in
@@ -411,7 +425,7 @@ let build_l2r_rew_scheme dep env ind kind =
mkApp (eq,[|lift 4 applied_ind;applied_sym_sym;mkRel 1|]),
applied_PR)),
mkApp (sym_involutive,
- Array.append (Context.Rel.to_extended_vect 3 mip.mind_arity_ctxt) [|mkVar varH|]),
+ Array.append (Context.Rel.to_extended_vect mkRel 3 mip.mind_arity_ctxt) [|mkVar varH|]),
[|main_body|])
else
main_body))))))
@@ -450,7 +464,7 @@ let build_l2r_forward_rew_scheme dep env ind kind =
get_sym_eq_data env indu in
let cstr n p =
mkApp (mkConstructUi(indu,1),
- Array.concat [Context.Rel.to_extended_vect n paramsctxt1;
+ Array.concat [Context.Rel.to_extended_vect mkRel n paramsctxt1;
rel_vect p nrealargs]) in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -458,12 +472,12 @@ let build_l2r_forward_rew_scheme dep env ind kind =
let applied_ind = build_dependent_inductive indu specif in
let applied_ind_P =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (4*nrealargs+2) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (4*nrealargs+2) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect (nrealargs+1) nrealargs]) in
let applied_ind_P' =
mkApp (mkIndU indu, Array.concat
- [Context.Rel.to_extended_vect (3*nrealargs+1) paramsctxt1;
+ [Context.Rel.to_extended_vect mkRel (3*nrealargs+1) paramsctxt1;
rel_vect 0 nrealargs;
rel_vect (2*nrealargs+1) nrealargs]) in
let realsign_P n = lift_rel_context (nrealargs*n+n) realsign in
@@ -541,7 +555,7 @@ let build_r2l_forward_rew_scheme dep env ind kind =
let ((mib,mip as specif),constrargs,realsign,paramsctxt,nrealargs) =
get_non_sym_eq_data env indu in
let cstr n =
- mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect n mib.mind_params_ctxt) in
+ mkApp (mkConstructUi(indu,1),Context.Rel.to_extended_vect mkRel n mib.mind_params_ctxt) in
let constrargs_cstr = constrargs@[cstr 0] in
let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in
let varHC = fresh env (Id.of_string "HC") in
@@ -557,8 +571,8 @@ let build_r2l_forward_rew_scheme dep env ind kind =
applist (mkVar varP,if dep then constrargs_cstr else constrargs) in
let applied_PG =
mkApp (mkVar varP,
- if dep then Context.Rel.to_extended_vect 0 realsign_ind
- else Context.Rel.to_extended_vect 1 realsign) in
+ if dep then Context.Rel.to_extended_vect mkRel 0 realsign_ind
+ else Context.Rel.to_extended_vect mkRel 1 realsign) in
let c =
(my_it_mkLambda_or_LetIn paramsctxt
(my_it_mkLambda_or_LetIn_name realsign_ind
@@ -596,7 +610,8 @@ let build_r2l_forward_rew_scheme dep env ind kind =
(**********************************************************************)
let fix_r2l_forward_rew_scheme (c, ctx') =
- let t = Retyping.get_type_of (Global.env()) Evd.empty c in
+ let t = Retyping.get_type_of (Global.env()) Evd.empty (EConstr.of_constr c) in
+ let t = EConstr.Unsafe.to_constr t in
let ctx,_ = decompose_prod_assum t in
match ctx with
| hp :: p :: ind :: indargs ->
@@ -605,9 +620,9 @@ let fix_r2l_forward_rew_scheme (c, ctx') =
(mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 1) p)
(mkLambda_or_LetIn (RelDecl.map_constr (liftn (-1) 2) hp)
(mkLambda_or_LetIn (RelDecl.map_constr (lift 2) ind)
- (Reductionops.whd_beta Evd.empty
- (applist (c,
- Context.Rel.to_extended_list 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))
+ (EConstr.Unsafe.to_constr (Reductionops.whd_beta Evd.empty
+ (EConstr.of_constr (applist (c,
+ Context.Rel.to_extended_list mkRel 3 indargs @ [mkRel 1;mkRel 3;mkRel 2]))))))))
in c', ctx'
| _ -> anomaly (Pp.str "Ill-formed non-dependent left-to-right rewriting scheme")
@@ -762,8 +777,8 @@ let build_congr env (eq,refl,ctx) ind =
(mkNamedLambda varH
(applist
(mkIndU indu,
- Context.Rel.to_extended_list (mip.mind_nrealargs+2) paramsctxt @
- Context.Rel.to_extended_list 0 realsign))
+ Context.Rel.to_extended_list mkRel (mip.mind_nrealargs+2) paramsctxt @
+ Context.Rel.to_extended_list mkRel 0 realsign))
(mkCase (ci,
my_it_mkLambda_or_LetIn_name
(lift_rel_context (mip.mind_nrealargs+3) realsign)
@@ -771,9 +786,9 @@ let build_congr env (eq,refl,ctx) ind =
(Anonymous,
applist
(mkIndU indu,
- Context.Rel.to_extended_list (2*mip.mind_nrealdecls+3)
+ Context.Rel.to_extended_list mkRel (2*mip.mind_nrealdecls+3)
paramsctxt
- @ Context.Rel.to_extended_list 0 realsign),
+ @ Context.Rel.to_extended_list mkRel 0 realsign),
mkApp (eq,
[|mkVar varB;
mkApp (mkVar varf, [|lift (2*mip.mind_nrealdecls+4) b|]);
diff --git a/tactics/equality.ml b/tactics/equality.ml
index d44dcf10d..7ae7446c8 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -6,18 +6,20 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Inductive
open Inductiveops
-open Environ
open Libnames
open Globnames
open Reductionops
@@ -144,7 +146,7 @@ let freeze_initial_evars sigma flags clause =
(* We take evars of the type: this may include old evars! For excluding *)
(* all old evars, including the ones occurring in the rewriting lemma, *)
(* we would have to take the clenv_value *)
- let newevars = Evd.evars_of_term (clenv_type clause) in
+ let newevars = Evarutil.undefined_evars_of_term sigma (clenv_type clause) in
let evars =
fold_undefined (fun evk _ evars ->
if Evar.Set.mem evk newevars then evars
@@ -165,8 +167,9 @@ let side_tac tac sidetac =
let instantiate_lemma_all frzevars gl c ty l l2r concl =
let env = Proofview.Goal.env gl in
+ let sigma = project gl in
let eqclause = pf_apply Clenv.make_clenv_binding gl (c,ty) l in
- let (equiv, args) = decompose_appvect (Clenv.clenv_type eqclause) in
+ let (equiv, args) = decompose_app_vect sigma (Clenv.clenv_type eqclause) in
let arglen = Array.length args in
let () = if arglen < 2 then error "The term provided is not an applied relation." in
let c1 = args.(arglen - 2) in
@@ -304,7 +307,7 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim =
else instantiate_lemma_all frzevars gl c t l l2r concl
in
let typ = match cls with
- | None -> pf_nf_concl gl
+ | None -> pf_concl gl
| Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl)
in
let cs = instantiate_lemma typ in
@@ -326,9 +329,9 @@ let (forward_general_setoid_rewrite_clause, general_setoid_rewrite_clause) = Hoo
let jmeq_same_dom gl = function
| None -> true (* already checked in Hipattern.find_eq_data_decompose *)
| Some t ->
- let rels, t = decompose_prod_assum t in
- let env = Environ.push_rel_context rels (Proofview.Goal.env gl) in
- match decompose_app t with
+ let rels, t = decompose_prod_assum (project gl) t in
+ let env = push_rel_context rels (Proofview.Goal.env gl) in
+ match decompose_app (project gl) t with
| _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2
| _ -> false
@@ -336,6 +339,8 @@ let jmeq_same_dom gl = function
eliminate lbeq on sort_of_gl. *)
let find_elim hdcncl lft2rgt dep cls ot gl =
+ let sigma = project gl in
+ let is_global gr c = Termops.is_global sigma gr c in
let inccl = Option.is_empty cls in
if (is_global Coqlib.glob_eq hdcncl ||
(is_global Coqlib.glob_jmeq hdcncl &&
@@ -343,7 +348,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
|| Flags.version_less_or_equal Flags.V8_2
then
let c =
- match kind_of_term hdcncl with
+ match EConstr.kind sigma hdcncl with
| Ind (ind_sp,u) ->
let pr1 =
lookup_eliminator ind_sp (elimination_sort_of_clause cls gl)
@@ -371,6 +376,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
assert false
in
let Sigma (elim, sigma, p) = Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in
+ let elim = EConstr.of_constr elim in
Sigma ((elim, Safe_typing.empty_private_constants), sigma, p)
else
let scheme_name = match dep, lft2rgt, inccl with
@@ -385,13 +391,14 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
| true, _, true -> rew_r2l_dep_scheme_kind
| true, _, false -> rew_r2l_forward_dep_scheme_kind
in
- match kind_of_term hdcncl with
+ match EConstr.kind sigma hdcncl with
| Ind (ind,u) ->
let c, eff = find_scheme scheme_name ind in
(* MS: cannot use pf_constr_of_global as the eliminator might be generated by side-effect *)
let Sigma (elim, sigma, p) =
Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c)
in
+ let elim = EConstr.of_constr elim in
Sigma ((elim, eff), sigma, p)
| _ -> assert false
@@ -400,12 +407,12 @@ let type_of_clause cls gl = match cls with
| Some id -> pf_get_hyp_typ id gl
let leibniz_rewrite_ebindings_clause cls lft2rgt tac c t l with_evars frzevars dep_proof_ok hdcncl =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in
- let isatomic = isProd (whd_zeta evd hdcncl) in
+ let isatomic = isProd evd (whd_zeta evd hdcncl) in
let dep_fun = if isatomic then dependent else dependent_no_evar in
let type_of_cls = type_of_clause cls gl in
- let dep = dep_proof_ok && dep_fun c type_of_cls in
+ let dep = dep_proof_ok && dep_fun evd c type_of_cls in
let Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in
let tac =
Proofview.tclEFFECTS effs <*>
@@ -441,8 +448,8 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
let ctype = get_type_of env sigma c in
- let rels, t = decompose_prod_assum (whd_betaiotazeta sigma ctype) in
- match match_with_equality_type t with
+ let rels, t = decompose_prod_assum sigma (whd_betaiotazeta sigma ctype) in
+ match match_with_equality_type sigma t with
| Some (hdcncl,args) -> (* Fast path: direct leibniz-like rewrite *)
let lft2rgt = adjust_rewriting_direction args lft2rgt in
leibniz_rewrite_ebindings_clause cls lft2rgt tac c (it_mkProd_or_LetIn t rels)
@@ -455,9 +462,10 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
end
begin function
| (e, info) ->
+ Proofview.tclEVARMAP >>= fun sigma ->
let env' = push_rel_context rels env in
let rels',t' = splay_prod_assum env' sigma t in (* Search for underlying eq *)
- match match_with_equality_type t' with
+ match match_with_equality_type sigma t' with
| Some (hdcncl,args) ->
let lft2rgt = adjust_rewriting_direction args lft2rgt in
leibniz_rewrite_ebindings_clause cls lft2rgt tac c
@@ -523,7 +531,7 @@ let general_rewrite_clause l2r with_evars ?tac c cl =
let do_hyps =
(* If the term to rewrite uses an hypothesis H, don't rewrite in H *)
let ids gl =
- let ids_in_c = Environ.global_vars_set (Global.env()) (fst c) in
+ let ids_in_c = Termops.global_vars_set (Global.env()) (project gl) (fst c) in
let ids_of_hyps = pf_ids_of_hyps gl in
Id.Set.fold (fun id l -> List.remove Id.equal id l) ids_in_c ids_of_hyps
in
@@ -713,7 +721,6 @@ let _ =
optread = (fun () -> !keep_proof_equalities_for_injection) ;
optwrite = (fun b -> keep_proof_equalities_for_injection := b) }
-
let find_positions env sigma t1 t2 =
let project env sorts posn t1 t2 =
let ty1 = get_type_of env sigma t1 in
@@ -724,7 +731,7 @@ let find_positions env sigma t1 t2 =
let rec findrec sorts posn t1 t2 =
let hd1,args1 = whd_all_stack env sigma t1 in
let hd2,args2 = whd_all_stack env sigma t2 in
- match (kind_of_term hd1, kind_of_term hd2) with
+ match (EConstr.kind sigma hd1, EConstr.kind sigma hd2) with
| Construct ((ind1,i1 as sp1),u1), Construct (sp2,_)
when Int.equal (List.length args1) (constructor_nallargs_env env sp1)
->
@@ -738,8 +745,10 @@ let find_positions env sigma t1 t2 =
let params1,rargs1 = List.chop nparams args1 in
let _,rargs2 = List.chop nparams args2 in
let (mib,mip) = lookup_mind_specif env ind1 in
+ let params1 = List.map EConstr.Unsafe.to_constr params1 in
+ let u1 = EInstance.kind sigma u1 in
let ctxt = (get_constructor ((ind1,u1),mib,mip,params1) i1).cs_args in
- let adjust i = Vars.adjust_rel_to_rel_context ctxt (i+1) - 1 in
+ let adjust i = CVars.adjust_rel_to_rel_context ctxt (i+1) - 1 in
List.flatten
(List.map2_i (fun i -> findrec sorts' ((sp1,adjust i)::posn))
0 rargs1 rargs2)
@@ -850,21 +859,23 @@ let descend_then env sigma head dirn =
let (mib,mip) = lookup_mind_specif env ind in
let cstr = get_constructors env indf in
let dirn_nlams = cstr.(dirn-1).cs_nargs in
- let dirn_env = push_rel_context cstr.(dirn-1).cs_args env in
+ let dirn_env = Environ.push_rel_context cstr.(dirn-1).cs_args env in
(dirn_nlams,
dirn_env,
(fun dirnval (dfltval,resty) ->
- let deparsign = make_arity_signature env true indf in
+ let deparsign = make_arity_signature env sigma true indf in
let p =
it_mkLambda_or_LetIn (lift (mip.mind_nrealargs+1) resty) deparsign in
let build_branch i =
let result = if Int.equal i dirn then dirnval else dfltval in
- it_mkLambda_or_LetIn_name env result cstr.(i-1).cs_args in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.(i-1).cs_args in
+ let args = name_context env sigma cs_args in
+ it_mkLambda_or_LetIn result args in
let brl =
List.map build_branch
(List.interval 1 (Array.length mip.mind_consnames)) in
let ci = make_case_info env ind RegularStyle in
- Inductiveops.make_case_or_project env indf ci p head (Array.of_list brl)))
+ Inductiveops.make_case_or_project env sigma indf ci p head (Array.of_list brl)))
(* Now we need to construct the discriminator, given a discriminable
position. This boils down to:
@@ -899,17 +910,22 @@ let build_selector env sigma dirn c ind special default =
let ind, _ = check_privacy env indp in
let typ = Retyping.get_type_of env sigma default in
let (mib,mip) = lookup_mind_specif env ind in
- let deparsign = make_arity_signature env true indf in
+ let deparsign = make_arity_signature env sigma true indf in
let p = it_mkLambda_or_LetIn typ deparsign in
let cstrs = get_constructors env indf in
let build_branch i =
let endpt = if Int.equal i dirn then special else default in
- it_mkLambda_or_LetIn endpt cstrs.(i-1).cs_args in
+ let args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstrs.(i-1).cs_args in
+ it_mkLambda_or_LetIn endpt args in
let brl =
List.map build_branch(List.interval 1 (Array.length mip.mind_consnames)) in
let ci = make_case_info env ind RegularStyle in
mkCase (ci, p, c, Array.of_list brl)
+let build_coq_False () = EConstr.of_constr (build_coq_False ())
+let build_coq_True () = EConstr.of_constr (build_coq_True ())
+let build_coq_I () = EConstr.of_constr (build_coq_I ())
+
let rec build_discriminator env sigma dirn c = function
| [] ->
let ind = get_type_of env sigma c in
@@ -935,9 +951,9 @@ let rec build_discriminator env sigma dirn c = function
let gen_absurdity id =
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
let hyp_typ = pf_get_hyp_typ id (Proofview.Goal.assume gl) in
- let hyp_typ = pf_nf_evar gl hyp_typ in
- if is_empty_type hyp_typ
+ if is_empty_type sigma hyp_typ
then
simplest_elim (mkVar id)
else
@@ -967,6 +983,7 @@ let discrimination_pf env sigma e (t,t1,t2) discriminator lbeq =
let absurd_term = build_coq_False () in
let eq_elim, eff = ind_scheme_of_eq lbeq in
let sigma, eq_elim = Evd.fresh_global (Global.env ()) sigma eq_elim in
+ let eq_elim = EConstr.of_constr eq_elim in
sigma, (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term),
eff
@@ -976,7 +993,7 @@ let apply_on_clause (f,t) clause =
let sigma = clause.evd in
let f_clause = mk_clenv_from_env clause.env sigma None (f,t) in
let argmv =
- (match kind_of_term (last_arg f_clause.templval.Evd.rebus) with
+ (match EConstr.kind sigma (last_arg f_clause.evd f_clause.templval.Evd.rebus) with
| Meta mv -> mv
| _ -> user_err (str "Ill-formed clause applicator.")) in
clenv_fchain ~with_univs:false argmv f_clause clause
@@ -998,7 +1015,7 @@ let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn =
let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause =
let sigma = eq_clause.evd in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
match find_positions env sigma t1 t2 with
| Inr _ ->
@@ -1008,7 +1025,7 @@ let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause =
end }
let onEquality with_evars tac (c,lbindc) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let type_of = pf_unsafe_type_of gl in
let reduce_to_quantified_ind = pf_apply Tacred.reduce_to_quantified_ind gl in
let t = type_of c in
@@ -1023,12 +1040,12 @@ let onEquality with_evars tac (c,lbindc) =
end }
let onNegatedEquality with_evars tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = Tacmach.New.project gl in
let ccl = Proofview.Goal.concl gl in
let env = Proofview.Goal.env gl in
- match kind_of_term (hnf_constr env sigma ccl) with
- | Prod (_,t,u) when is_empty_type u ->
+ match EConstr.kind sigma (hnf_constr env sigma ccl) with
+ | Prod (_,t,u) when is_empty_type sigma u ->
tclTHEN introf
(onLastHypId (fun id ->
onEquality with_evars tac (mkVar id,NoBindings)))
@@ -1082,7 +1099,7 @@ let find_sigma_data env s = build_sigma_type ()
*)
let make_tuple env sigma (rterm,rty) lind =
- assert (dependent (mkRel lind) rty);
+ assert (not (noccurn sigma lind rty));
let sigdata = find_sigma_data env (get_sort_of env sigma rty) in
let sigma, a = type_of ~refresh:true env sigma (mkRel lind) in
let na = Context.Rel.Declaration.get_name (lookup_rel lind env) in
@@ -1092,6 +1109,8 @@ let make_tuple env sigma (rterm,rty) lind =
let p = mkLambda (na, a, rty) in
let sigma, exist_term = Evd.fresh_global env sigma sigdata.intro in
let sigma, sig_term = Evd.fresh_global env sigma sigdata.typ in
+ let exist_term = EConstr.of_constr exist_term in
+ let sig_term = EConstr.of_constr sig_term in
sigma,
(applist(exist_term,[a;p;(mkRel lind);rterm]),
applist(sig_term,[a;p]))
@@ -1104,9 +1123,9 @@ let make_tuple env sigma (rterm,rty) lind =
normalization *)
let minimal_free_rels env sigma (c,cty) =
- let cty_rels = free_rels cty in
+ let cty_rels = free_rels sigma cty in
let cty' = simpl env sigma cty in
- let rels' = free_rels cty' in
+ let rels' = free_rels sigma cty' in
if Int.Set.subset cty_rels rels' then
(cty,cty_rels)
else
@@ -1174,19 +1193,18 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt =
error "Cannot solve a unification problem."
else
let (a,p_i_minus_1) = match whd_beta_stack !evdref p_i with
- | (_sigS,[a;p]) -> (a,p)
+ | (_sigS,[a;p]) -> (a, p)
| _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in
let ev = Evarutil.e_new_evar env evdref a in
- let rty = beta_applist(p_i_minus_1,[ev]) in
+ let rty = beta_applist sigma (p_i_minus_1,[ev]) in
let tuple_tail = sigrec_clausal_form (siglen-1) rty in
- match
- Evd.existential_opt_value !evdref
- (destEvar ev)
- with
+ let evopt = match EConstr.kind !evdref ev with Evar _ -> None | _ -> Some ev in
+ match evopt with
| Some w ->
- let w_type = unsafe_type_of env sigma w in
+ let w_type = unsafe_type_of env !evdref w in
if Evarconv.e_cumul env evdref w_type a then
let exist_term = Evarutil.evd_comb1 (Evd.fresh_global env) evdref sigdata.intro in
+ let exist_term = EConstr.of_constr exist_term in
applist(exist_term,[a;p_i_minus_1;w;tuple_tail])
else
error "Cannot solve a unification problem."
@@ -1267,7 +1285,7 @@ let make_iterated_tuple env sigma dflt (z,zty) =
let sigma, (tuple,tuplety) =
List.fold_left (fun (sigma, t) -> make_tuple env sigma t) (sigma, (z,zty)) sorted_rels
in
- assert (closed0 tuplety);
+ assert (closed0 sigma tuplety);
let n = List.length sorted_rels in
let sigma, dfltval = sig_clausal_form env sigma sort_of_zty n tuplety dflt in
sigma, (tuple,tuplety,dfltval)
@@ -1287,56 +1305,46 @@ let build_injector env sigma dflt c cpath =
let sigma, (injcode,resty,_) = build_injrec env sigma dflt c cpath in
sigma, (injcode,resty)
-(*
-let try_delta_expand env sigma t =
- let whdt = whd_all env sigma t in
- let rec hd_rec c =
- match kind_of_term c with
- | Construct _ -> whdt
- | App (f,_) -> hd_rec f
- | Cast (c,_,_) -> hd_rec c
- | _ -> t
- in
- hd_rec whdt
-*)
-
let eq_dec_scheme_kind_name = ref (fun _ -> failwith "eq_dec_scheme undefined")
let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k)
let inject_if_homogenous_dependent_pair ty =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
try
+ let sigma = Tacmach.New.project gl in
let eq,u,(t,t1,t2) = find_this_eq_data_decompose gl ty in
(* fetch the informations of the pair *)
let ceq = Universes.constr_of_global Coqlib.glob_eq in
+ let ceq = EConstr.of_constr ceq in
let sigTconstr () = (Coqlib.build_sigma_type()).Coqlib.typ in
let existTconstr () = (Coqlib.build_sigma_type()).Coqlib.intro in
(* check whether the equality deals with dep pairs or not *)
- let eqTypeDest = fst (decompose_app t) in
- if not (Globnames.is_global (sigTconstr()) eqTypeDest) then raise Exit;
- let hd1,ar1 = decompose_app_vect t1 and
- hd2,ar2 = decompose_app_vect t2 in
- if not (Globnames.is_global (existTconstr()) hd1) then raise Exit;
- if not (Globnames.is_global (existTconstr()) hd2) then raise Exit;
- let ind,_ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in
+ let eqTypeDest = fst (decompose_app sigma t) in
+ if not (Termops.is_global sigma (sigTconstr()) eqTypeDest) then raise Exit;
+ let hd1,ar1 = decompose_app_vect sigma t1 and
+ hd2,ar2 = decompose_app_vect sigma t2 in
+ if not (Termops.is_global sigma (existTconstr()) hd1) then raise Exit;
+ if not (Termops.is_global sigma (existTconstr()) hd2) then raise Exit;
+ let (ind, _), _ = try pf_apply find_mrectype gl ar1.(0) with Not_found -> raise Exit in
(* check if the user has declared the dec principle *)
(* and compare the fst arguments of the dep pair *)
(* Note: should work even if not an inductive type, but the table only *)
(* knows inductive types *)
- if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) (fst ind) &&
+ if not (Ind_tables.check_scheme (!eq_dec_scheme_kind_name()) ind &&
pf_apply is_conv gl ar1.(2) ar2.(2)) then raise Exit;
Coqlib.check_required_library ["Coq";"Logic";"Eqdep_dec"];
let new_eq_args = [|pf_unsafe_type_of gl ar1.(3);ar1.(3);ar2.(3)|] in
let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing"
["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in
- let c, eff = find_scheme (!eq_dec_scheme_kind_name()) (Univ.out_punivs ind) in
+ let inj2 = EConstr.of_constr inj2 in
+ let c, eff = find_scheme (!eq_dec_scheme_kind_name()) ind in
(* cut with the good equality and prove the requested goal *)
tclTHENLIST
[Proofview.tclEFFECTS eff;
intro;
onLastHyp (fun hyp ->
tclTHENS (cut (mkApp (ceq,new_eq_args)))
- [clear [destVar hyp];
+ [clear [destVar sigma hyp];
Proofview.V82.tactic (Tacmach.refine
(mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))
])]
@@ -1350,7 +1358,7 @@ let inject_if_homogenous_dependent_pair ty =
let simplify_args env sigma t =
(* Quick hack to reduce in arguments of eq only *)
- match decompose_app t with
+ match decompose_app sigma t with
| eq, [t;c1;c2] -> applist (eq,[t;simpl env sigma c1;simpl env sigma c2])
| eq, [t1;c1;t2;c2] -> applist (eq,[t1;simpl env sigma c1;t2;simpl env sigma c2])
| _ -> t
@@ -1365,6 +1373,7 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
let sigma, (injbody,resty) = build_injector e_env !evdref t1' (mkVar e) cpath in
let injfun = mkNamedLambda e t injbody in
let sigma,congr = Evd.fresh_global env sigma eq.congr in
+ let congr = EConstr.of_constr congr in
let pf = applist(congr,[t;resty;injfun;t1;t2]) in
let sigma, pf_typ = Typing.type_of env sigma pf in
let inj_clause = apply_on_clause (pf,pf_typ) eq_clause in
@@ -1426,7 +1435,8 @@ let injEq ?(old=false) with_evars clear_flag ipats =
match ipats_style with
| Some ipats ->
Proofview.Goal.enter { enter = begin fun gl ->
- let destopt = match kind_of_term c with
+ let sigma = project gl in
+ let destopt = match EConstr.kind sigma c with
| Var id -> get_previous_hyp_position id gl
| _ -> MoveLast in
let clear_tac =
@@ -1455,7 +1465,7 @@ let injConcl = injClause None false None
let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id)))
let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma = clause.evd in
let env = Proofview.Goal.env gl in
match find_positions env sigma t1 t2 with
@@ -1476,9 +1486,9 @@ let dEq with_evars =
dEqThen with_evars (fun clear_flag c x ->
(apply_clear_request clear_flag (use_clear_hyp_by_default ()) c))
-let intro_decomp_eq tac data cl =
+let intro_decomp_eq tac data (c, t) =
Proofview.Goal.enter { enter = begin fun gl ->
- let cl = pf_apply make_clenv_binding gl cl NoBindings in
+ let cl = pf_apply make_clenv_binding gl (c, t) NoBindings in
decompEqThen (fun _ -> tac) data cl
end }
@@ -1516,14 +1526,14 @@ let _ = declare_intro_decomp_eq intro_decomp_eq
*)
-let decomp_tuple_term env c t =
+let decomp_tuple_term env sigma c t =
let rec decomprec inner_code ex exty =
let iterated_decomp =
try
- let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose ex in
+ let ({proj1=p1; proj2=p2}),(i,a,p,car,cdr) = find_sigma_data_decompose sigma ex in
let car_code = applist (mkConstU (destConstRef p1,i),[a;p;inner_code])
and cdr_code = applist (mkConstU (destConstRef p2,i),[a;p;inner_code]) in
- let cdrtyp = beta_applist (p,[car]) in
+ let cdrtyp = beta_applist sigma (p,[car]) in
List.map (fun l -> ((car,a),car_code)::l) (decomprec cdr_code cdr cdrtyp)
with Constr_matching.PatternMatchingFailure ->
[]
@@ -1534,8 +1544,8 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
let sigma = Sigma.to_evar_map sigma in
let typ = get_type_of env sigma dep_pair1 in
(* We find all possible decompositions *)
- let decomps1 = decomp_tuple_term env dep_pair1 typ in
- let decomps2 = decomp_tuple_term env dep_pair2 typ in
+ let decomps1 = decomp_tuple_term env sigma dep_pair1 typ in
+ let decomps2 = decomp_tuple_term env sigma dep_pair2 typ in
(* We adjust to the shortest decomposition *)
let n = min (List.length decomps1) (List.length decomps2) in
let decomp1 = List.nth decomps1 (n-1) in
@@ -1547,10 +1557,10 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
(* We build the expected goal *)
let abst_B =
List.fold_right
- (fun (e,t) body -> lambda_create env (t,subst_term e body)) e1_list b in
- let pred_body = beta_applist(abst_B,proj_list) in
- let body = mkApp (lambda_create env (typ,pred_body),[|dep_pair1|]) in
- let expected_goal = beta_applist (abst_B,List.map fst e2_list) in
+ (fun (e,t) body -> lambda_create env sigma (t,subst_term sigma e body)) e1_list b in
+ let pred_body = beta_applist sigma (abst_B,proj_list) in
+ let body = mkApp (lambda_create env sigma (typ,pred_body),[|dep_pair1|]) in
+ let expected_goal = beta_applist sigma (abst_B,List.map fst e2_list) in
(* Simulate now the normalisation treatment made by Logic.mk_refgoals *)
let expected_goal = nf_betaiota sigma expected_goal in
(* Retype to get universes right *)
@@ -1564,7 +1574,7 @@ let subst_tuple_term env sigma dep_pair1 dep_pair2 b =
(* on for further iterated sigma-tuples *)
let cutSubstInConcl l2r eqn =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in
@@ -1583,7 +1593,7 @@ let cutSubstInConcl l2r eqn =
end }
let cutSubstInHyp l2r eqn id =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let (lbeq,u,(t,e1,e2)) = find_eq_data_decompose gl eqn in
@@ -1672,14 +1682,14 @@ exception FoundHyp of (Id.t * constr * bool)
let is_eq_x gl x d =
let id = NamedDecl.get_id d in
try
- let is_var id c = match kind_of_term c with
+ let is_var id c = match EConstr.kind (project gl) c with
| Var id' -> Id.equal id id'
| _ -> false
in
let c = pf_nf_evar gl (NamedDecl.get_type d) in
let (_,lhs,rhs) = pi3 (find_eq_data_decompose gl c) in
- if (is_var x lhs) && not (local_occur_var x rhs) then raise (FoundHyp (id,rhs,true));
- if (is_var x rhs) && not (local_occur_var x lhs) then raise (FoundHyp (id,lhs,false))
+ if (is_var x lhs) && not (local_occur_var (project gl) x rhs) then raise (FoundHyp (id,rhs,true));
+ if (is_var x rhs) && not (local_occur_var (project gl) x lhs) then raise (FoundHyp (id,lhs,false))
with Constr_matching.PatternMatchingFailure ->
()
@@ -1689,6 +1699,7 @@ let is_eq_x gl x d =
let subst_one dep_proof_ok x (hyp,rhs,dir) =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
(* The set of hypotheses using x *)
@@ -1696,7 +1707,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
List.rev (pi3 (List.fold_right (fun dcl (dest,deps,allhyps) ->
let id = NamedDecl.get_id dcl in
if not (Id.equal id hyp)
- && List.exists (fun y -> occur_var_in_decl env y dcl) deps
+ && List.exists (fun y -> occur_var_in_decl env sigma y dcl) deps
then
let id_dest = if !regular_subst_tactic then dest else MoveLast in
(dest,id::deps,(id_dest,id)::allhyps)
@@ -1705,7 +1716,7 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
hyps
(MoveBefore x,[x],[]))) in (* In practice, no dep hyps before x, so MoveBefore x is good enough *)
(* Decides if x appears in conclusion *)
- let depconcl = occur_var env x concl in
+ let depconcl = occur_var env sigma x concl in
let need_rewrite = not (List.is_empty dephyps) || depconcl in
tclTHENLIST
((if need_rewrite then
@@ -1768,13 +1779,15 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
let find_equations gl =
let gl = Proofview.Goal.assume gl in
let env = Proofview.Goal.env gl in
+ let sigma = project gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let select_equation_name decl =
try
let lbeq,u,(_,x,y) = find_eq_data_decompose (NamedDecl.get_type decl) in
+ let u = EInstance.kind sigma u in
let eq = Universes.constr_of_global_univ (lbeq.eq,u) in
if flags.only_leibniz then restrict_to_eq_and_identity eq;
- match kind_of_term x, kind_of_term y with
+ match EConstr.kind sigma x, EConstr.kind sigma y with
| Var z, _ when not (is_evaluable env (EvalVarRef z)) ->
Some (NamedDecl.get_id decl)
| _, Var z when not (is_evaluable env (EvalVarRef z)) ->
@@ -1791,22 +1804,23 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
let process hyp =
Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.assume gl in
+ let sigma = project gl in
let env = Proofview.Goal.env gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let c = pf_get_hyp hyp gl |> NamedDecl.get_type in
let _,_,(_,x,y) = find_eq_data_decompose c in
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
- if Term.eq_constr x y then Proofview.tclUNIT () else
- match kind_of_term x, kind_of_term y with
- | Var x', _ when not (occur_term x y) && not (is_evaluable env (EvalVarRef x')) ->
+ if EConstr.eq_constr sigma x y then Proofview.tclUNIT () else
+ match EConstr.kind sigma x, EConstr.kind sigma y with
+ | Var x', _ when not (occur_term sigma x y) && not (is_evaluable env (EvalVarRef x')) ->
subst_one flags.rewrite_dependent_proof x' (hyp,y,true)
- | _, Var y' when not (occur_term y x) && not (is_evaluable env (EvalVarRef y')) ->
+ | _, Var y' when not (occur_term sigma y x) && not (is_evaluable env (EvalVarRef y')) ->
subst_one flags.rewrite_dependent_proof y' (hyp,x,false)
| _ ->
Proofview.tclUNIT ()
end }
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let ids = find_equations gl in
tclMAP process ids
end }
@@ -1816,17 +1830,19 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
(* Old implementation, not able to manage configurations like a=b, a=t,
or situations like "a = S b, b = S a", or also accidentally unfolding
let-ins *)
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
let find_eq_data_decompose = find_eq_data_decompose gl in
let test (_,c) =
try
let lbeq,u,(_,x,y) = find_eq_data_decompose c in
+ let u = EInstance.kind sigma u in
let eq = Universes.constr_of_global_univ (lbeq.eq,u) in
if flags.only_leibniz then restrict_to_eq_and_identity eq;
(* J.F.: added to prevent failure on goal containing x=x as an hyp *)
- if Term.eq_constr x y then failwith "caught";
- match kind_of_term x with Var x -> x | _ ->
- match kind_of_term y with Var y -> y | _ -> failwith "caught"
+ if EConstr.eq_constr sigma x y then failwith "caught";
+ match EConstr.kind sigma x with Var x -> x | _ ->
+ match EConstr.kind sigma y with Var y -> y | _ -> failwith "caught"
with Constr_matching.PatternMatchingFailure -> failwith "caught" in
let test p = try Some (test p) with Failure _ -> None in
let hyps = pf_hyps_types gl in
@@ -1870,7 +1886,7 @@ let rewrite_assumption_cond cond_eq_term cl =
with | Failure _ | UserError _ -> arec rest gl
end
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let gl = Proofview.Goal.lift gl Sigma.Unsafe.le in
let hyps = Proofview.Goal.hyps gl in
arec hyps gl
diff --git a/tactics/equality.mli b/tactics/equality.mli
index 6a4a8126e..5467b4af2 100644
--- a/tactics/equality.mli
+++ b/tactics/equality.mli
@@ -10,6 +10,7 @@
open Names
open Term
open Evd
+open EConstr
open Environ
open Ind_tables
open Locus
diff --git a/tactics/hints.ml b/tactics/hints.ml
index 59d015fa2..77ed4330c 100644
--- a/tactics/hints.ml
+++ b/tactics/hints.ml
@@ -6,12 +6,16 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open Util
open CErrors
open Names
-open Vars
open Term
+open Evd
+open EConstr
+open Vars
open Environ
open Mod_subst
open Globnames
@@ -21,7 +25,6 @@ open Libnames
open Smartlocate
open Misctypes
open Tactypes
-open Evd
open Termops
open Inductiveops
open Typing
@@ -45,23 +48,26 @@ type debug = Debug | Info | Off
exception Bound
-let head_constr_bound t =
- let t = strip_outer_cast t in
- let _,ccl = decompose_prod_assum t in
- let hd,args = decompose_app ccl in
- match kind_of_term hd with
- | Const _ | Ind _ | Construct _ | Var _ -> hd
- | Proj (p, _) -> mkConst (Projection.constant p)
- | _ -> raise Bound
-
-let head_constr c =
- try head_constr_bound c with Bound -> error "Bound head variable."
-
-let decompose_app_bound t =
- let t = strip_outer_cast t in
- let _,ccl = decompose_prod_assum t in
- let hd,args = decompose_app_vect ccl in
- match kind_of_term hd with
+let head_constr_bound sigma t =
+ let t = strip_outer_cast sigma t in
+ let _,ccl = decompose_prod_assum sigma t in
+ let hd,args = decompose_app sigma ccl in
+ match EConstr.kind sigma hd with
+ | Const (c, _) -> ConstRef c
+ | Ind (i, _) -> IndRef i
+ | Construct (c, _) -> ConstructRef c
+ | Var id -> VarRef id
+ | Proj (p, _) -> ConstRef (Projection.constant p)
+ | _ -> raise Bound
+
+let head_constr sigma c =
+ try head_constr_bound sigma c with Bound -> error "Bound head variable."
+
+let decompose_app_bound sigma t =
+ let t = strip_outer_cast sigma t in
+ let _,ccl = decompose_prod_assum sigma t in
+ let hd,args = decompose_app_vect sigma ccl in
+ match EConstr.kind sigma hd with
| Const (c,u) -> ConstRef c, args
| Ind (i,u) -> IndRef i, args
| Construct (c,u) -> ConstructRef c, args
@@ -255,8 +261,8 @@ let rebuild_dn st se =
in
{ se with sentry_bnet = dn' }
-let lookup_tacs concl st se =
- let l' = Bounded_net.lookup st se.sentry_bnet concl in
+let lookup_tacs sigma concl st se =
+ let l' = Bounded_net.lookup sigma st se.sentry_bnet concl in
let sl' = List.stable_sort pri_order_int l' in
List.merge pri_order_int se.sentry_nopat sl'
@@ -267,10 +273,10 @@ let is_transparent_gr (ids, csts) = function
| ConstRef cst -> Cpred.mem cst csts
| IndRef _ | ConstructRef _ -> false
-let strip_params env c =
- match kind_of_term c with
+let strip_params env sigma c =
+ match EConstr.kind sigma c with
| App (f, args) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Const (p,_) ->
let cb = lookup_constant p env in
(match cb.Declarations.const_proj with
@@ -289,7 +295,7 @@ let instantiate_hint env sigma p =
let sigma = Evd.merge_context_set univ_flexible sigma ctx in
let cl = mk_clenv_from_env env sigma None (c,cty) in
{cl with templval =
- { cl.templval with rebus = strip_params env cl.templval.rebus };
+ { cl.templval with rebus = strip_params env sigma cl.templval.rebus };
env = empty_env}
in
let code = match p.code.obj with
@@ -473,11 +479,11 @@ val empty : ?name:hint_db_name -> transparent_state -> bool -> t
val find : global_reference -> t -> search_entry
val map_none : secvars:Id.Pred.t -> t -> full_hint list
val map_all : secvars:Id.Pred.t -> global_reference -> t -> full_hint list
-val map_existential : secvars:Id.Pred.t ->
+val map_existential : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
-val map_eauto : secvars:Id.Pred.t ->
+val map_eauto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
-val map_auto : secvars:Id.Pred.t ->
+val map_auto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
val add_one : env -> evar_map -> hint_entry -> t -> t
val add_list : env -> evar_map -> hint_entry list -> t -> t
@@ -533,21 +539,32 @@ struct
(** Warn about no longer typable hint? *)
None
- let match_mode m arg =
+ let head_evar sigma c =
+ let rec hrec c = match EConstr.kind sigma c with
+ | Evar (evk,_) -> evk
+ | Case (_,_,c,_) -> hrec c
+ | App (c,_) -> hrec c
+ | Cast (c,_,_) -> hrec c
+ | Proj (p, c) -> hrec c
+ | _ -> raise Evarutil.NoHeadEvar
+ in
+ hrec c
+
+ let match_mode sigma m arg =
match m with
- | ModeInput -> not (occur_existential arg)
+ | ModeInput -> not (occur_existential sigma arg)
| ModeNoHeadEvar ->
- Evarutil.(try ignore(head_evar arg); false
- with NoHeadEvar -> true)
+ (try ignore(head_evar sigma arg); false
+ with Evarutil.NoHeadEvar -> true)
| ModeOutput -> true
- let matches_mode args mode =
+ let matches_mode sigma args mode =
Array.length mode == Array.length args &&
- Array.for_all2 match_mode mode args
+ Array.for_all2 (match_mode sigma) mode args
- let matches_modes args modes =
+ let matches_modes sigma args modes =
if List.is_empty modes then true
- else List.exists (matches_mode args) modes
+ else List.exists (matches_mode sigma args) modes
let merge_entry secvars db nopat pat =
let h = List.sort pri_order_int (List.map snd db.hintdb_nopat) in
@@ -563,24 +580,24 @@ struct
merge_entry secvars db se.sentry_nopat se.sentry_pat
(** Precondition: concl has no existentials *)
- let map_auto ~secvars (k,args) concl db =
+ let map_auto sigma ~secvars (k,args) concl db =
let se = find k db in
let st = if db.use_dn then (Some db.hintdb_state) else None in
- let pat = lookup_tacs concl st se in
+ let pat = lookup_tacs sigma concl st se in
merge_entry secvars db [] pat
- let map_existential ~secvars (k,args) concl db =
+ let map_existential sigma ~secvars (k,args) concl db =
let se = find k db in
- if matches_modes args se.sentry_mode then
+ if matches_modes sigma args se.sentry_mode then
merge_entry secvars db se.sentry_nopat se.sentry_pat
else merge_entry secvars db [] []
(* [c] contains an existential *)
- let map_eauto ~secvars (k,args) concl db =
+ let map_eauto sigma ~secvars (k,args) concl db =
let se = find k db in
- if matches_modes args se.sentry_mode then
+ if matches_modes sigma args se.sentry_mode then
let st = if db.use_dn then Some db.hintdb_state else None in
- let pat = lookup_tacs concl st se in
+ let pat = lookup_tacs sigma concl st se in
merge_entry secvars db [] pat
else merge_entry secvars db [] []
@@ -745,8 +762,8 @@ let _ = Summary.declare_summary "search"
(* Auxiliary functions to prepare AUTOHINT objects *)
(**************************************************************************)
-let rec nb_hyp c = match kind_of_term c with
- | Prod(_,_,c2) -> if noccurn 1 c2 then 1+(nb_hyp c2) else nb_hyp c2
+let rec nb_hyp sigma c = match EConstr.kind sigma c with
+ | Prod(_,_,c2) -> if noccurn sigma 1 c2 then 1+(nb_hyp sigma c2) else nb_hyp sigma c2
| _ -> 0
(* adding and removing tactics in the search table *)
@@ -763,19 +780,19 @@ let secvars_of_idset s =
Id.Pred.add id p
else p) s Id.Pred.empty
-let secvars_of_constr env c =
- secvars_of_idset (global_vars_set env c)
+let secvars_of_constr env sigma c =
+ secvars_of_idset (Termops.global_vars_set env sigma c)
let secvars_of_global env gr =
secvars_of_idset (vars_of_global_reference env gr)
let make_exact_entry env sigma info poly ?(name=PathAny) (c, cty, ctx) =
- let secvars = secvars_of_constr env c in
- let cty = strip_outer_cast cty in
- match kind_of_term cty with
+ let secvars = secvars_of_constr env sigma c in
+ let cty = strip_outer_cast sigma cty in
+ match EConstr.kind sigma cty with
| Prod _ -> failwith "make_exact_entry"
| _ ->
- let pat = Patternops.pattern_of_constr env sigma cty in
+ let pat = Patternops.pattern_of_constr env sigma (EConstr.to_constr sigma cty) in
let hd =
try head_pattern_bound pat
with BoundPattern -> failwith "make_exact_entry"
@@ -792,18 +809,18 @@ let make_exact_entry env sigma info poly ?(name=PathAny) (c, cty, ctx) =
let make_apply_entry env sigma (eapply,hnf,verbose) info poly ?(name=PathAny) (c, cty, ctx) =
let cty = if hnf then hnf_constr env sigma cty else cty in
- match kind_of_term cty with
+ match EConstr.kind sigma cty with
| Prod _ ->
let sigma' = Evd.merge_context_set univ_flexible sigma ctx in
let ce = mk_clenv_from_env env sigma' None (c,cty) in
let c' = clenv_type (* ~reduce:false *) ce in
- let pat = Patternops.pattern_of_constr env ce.evd c' in
+ let pat = Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma c') in
let hd =
try head_pattern_bound pat
with BoundPattern -> failwith "make_apply_entry" in
let nmiss = List.length (clenv_missing ce) in
- let secvars = secvars_of_constr env c in
- let pri = match info.hint_priority with None -> nb_hyp cty + nmiss | Some p -> p in
+ let secvars = secvars_of_constr env sigma c in
+ let pri = match info.hint_priority with None -> nb_hyp sigma' cty + nmiss | Some p -> p in
let pat = match info.hint_pattern with
| Some p -> snd p | None -> pat
in
@@ -816,7 +833,7 @@ let make_apply_entry env sigma (eapply,hnf,verbose) info poly ?(name=PathAny) (c
else begin
if not eapply then failwith "make_apply_entry";
if verbose then
- Feedback.msg_info (str "the hint: eapply " ++ pr_lconstr c ++
+ Feedback.msg_info (str "the hint: eapply " ++ pr_leconstr_env env sigma' c ++
str " will only be used by eauto");
(Some hd,
{ pri; poly; pat = Some pat; name;
@@ -833,7 +850,7 @@ let pr_hint_term env sigma ctx = function
| IsGlobRef gr -> pr_global gr
| IsConstr (c, ctx) ->
let sigma = Evd.merge_context_set Evd.univ_flexible sigma ctx in
- pr_constr_env env sigma c
+ pr_econstr_env env sigma c
(** We need an object to record the side-effect of registering
global universes associated with a hint. *)
@@ -859,7 +876,8 @@ let fresh_global_or_constr env sigma poly cr =
let isgr, (c, ctx) =
match cr with
| IsGlobRef gr ->
- true, Universes.fresh_global_instance env gr
+ let (c, ctx) = Universes.fresh_global_instance env gr in
+ true, (EConstr.of_constr c, ctx)
| IsConstr (c, ctx) -> false, (c, ctx)
in
if poly then (c, ctx)
@@ -882,7 +900,7 @@ let make_resolves env sigma flags info poly ?name cr =
in
if List.is_empty ents then
user_err ~hdr:"Hint"
- (pr_lconstr c ++ spc() ++
+ (pr_leconstr_env env sigma c ++ spc() ++
(if pi1 flags then str"cannot be used as a hint."
else str "can be used as a hint only for eauto."));
ents
@@ -924,6 +942,7 @@ let make_extern pri pat tacast =
code = with_uid (Extern tacast) })
let make_mode ref m =
+ let open Term in
let ty = Global.type_of_global_unsafe ref in
let ctx, t = decompose_prod ty in
let n = List.length ctx in
@@ -938,14 +957,14 @@ let make_trivial env sigma poly ?(name=PathAny) r =
let c,ctx = fresh_global_or_constr env sigma poly r in
let sigma = Evd.merge_context_set univ_flexible sigma ctx in
let t = hnf_constr env sigma (unsafe_type_of env sigma c) in
- let hd = head_of_constr_reference (head_constr t) in
+ let hd = head_constr sigma t in
let ce = mk_clenv_from_env env sigma None (c,t) in
(Some hd, { pri=1;
poly = poly;
- pat = Some (Patternops.pattern_of_constr env ce.evd (clenv_type ce));
+ pat = Some (Patternops.pattern_of_constr env ce.evd (EConstr.to_constr sigma (clenv_type ce)));
name = name;
db = None;
- secvars = secvars_of_constr env c;
+ secvars = secvars_of_constr env sigma c;
code= with_uid (Res_pf_THEN_trivial_fail(c,t,ctx)) })
@@ -1039,14 +1058,16 @@ let cache_autohint (kn, obj) =
let subst_autohint (subst, obj) =
let subst_key gr =
let (lab'', elab') = subst_global subst gr in
+ let elab' = EConstr.of_constr elab' in
let gr' =
- (try head_of_constr_reference (head_constr_bound elab')
+ (try head_constr_bound Evd.empty elab'
with Bound -> lab'')
in if gr' == gr then gr else gr'
in
let subst_hint (k,data as hint) =
let k' = Option.smartmap subst_key k in
let pat' = Option.smartmap (subst_pattern subst) data.pat in
+ let subst_mps subst c = EConstr.of_constr (subst_mps subst (EConstr.Unsafe.to_constr c)) in
let code' = match data.code.obj with
| Res_pf (c,t,ctx) ->
let c' = subst_mps subst c in
@@ -1218,29 +1239,32 @@ let prepare_hint check (poly,local) env init (sigma,c) =
It is actually a bit stupid to generalize over evars since the first
thing make_resolves will do is to re-instantiate the products *)
let sigma, subst = Evd.nf_univ_variables sigma in
- let c = Vars.subst_univs_constr subst (Evarutil.nf_evar sigma c) in
- let c = drop_extra_implicit_args c in
- let vars = ref (collect_vars c) in
+ let c = Evarutil.nf_evar sigma c in
+ let c = EConstr.Unsafe.to_constr c in
+ let c = CVars.subst_univs_constr subst c in
+ let c = EConstr.of_constr c in
+ let c = drop_extra_implicit_args sigma c in
+ let vars = ref (collect_vars sigma c) in
let subst = ref [] in
- let rec find_next_evar c = match kind_of_term c with
+ let rec find_next_evar c = match EConstr.kind sigma c with
| Evar (evk,args as ev) ->
(* We skip the test whether args is the identity or not *)
- let t = Evarutil.nf_evar sigma (existential_type sigma ev) in
- let t = List.fold_right (fun (e,id) c -> replace_term e id c) !subst t in
- if not (closed0 c) then
+ let t = existential_type sigma ev in
+ let t = List.fold_right (fun (e,id) c -> replace_term sigma e id c) !subst t in
+ if not (closed0 sigma c) then
error "Hints with holes dependent on a bound variable not supported.";
- if occur_existential t then
+ if occur_existential sigma t then
(* Not clever enough to construct dependency graph of evars *)
error "Not clever enough to deal with evars dependent in other evars.";
raise (Found (c,t))
- | _ -> Constr.iter find_next_evar c in
+ | _ -> EConstr.iter sigma find_next_evar c in
let rec iter c =
try find_next_evar c; c
with Found (evar,t) ->
let id = next_ident_away_from default_prepare_hint_ident (fun id -> Id.Set.mem id !vars) in
vars := Id.Set.add id !vars;
subst := (evar,mkVar id)::!subst;
- mkNamedLambda id t (iter (replace_term evar (mkVar id) c)) in
+ mkNamedLambda id t (iter (replace_term sigma evar (mkVar id) c)) in
let c' = iter c in
if check then Pretyping.check_evars (Global.env()) Evd.empty sigma c';
let diff = Univ.ContextSet.diff (Evd.universe_context_set sigma) (Evd.universe_context_set init) in
@@ -1321,7 +1345,7 @@ let add_hints local dbnames0 h =
let expand_constructor_hints env sigma lems =
List.map_append (fun (evd,lem) ->
- match kind_of_term lem with
+ match EConstr.kind sigma lem with
| Ind (ind,u) ->
List.init (nconstructors ind)
(fun i ->
@@ -1351,7 +1375,7 @@ let make_local_hint_db env sigma ts eapply lems =
(Sigma.to_evar_map sigma, c)
in
let lems = List.map map lems in
- let sign = Environ.named_context env in
+ let sign = EConstr.named_context env in
let ts = match ts with
| None -> Hint_db.transparent_state (searchtable_map "core")
| Some ts -> ts
@@ -1376,7 +1400,7 @@ let make_db_list dbnames =
(* Functions for printing the hints *)
(**************************************************************************)
-let pr_hint_elt (c, _, _) = pr_constr c
+let pr_hint_elt (c, _, _) = pr_econstr c
let pr_hint h = match h.obj with
| Res_pf (c, _) -> (str"simple apply " ++ pr_hint_elt c)
@@ -1426,15 +1450,15 @@ let pr_hint_ref ref = pr_hint_list_for_head ref
(* Print all hints associated to head id in any database *)
-let pr_hint_term cl =
+let pr_hint_term sigma cl =
try
let dbs = current_db () in
let valid_dbs =
let fn = try
- let hdc = decompose_app_bound cl in
- if occur_existential cl then
- Hint_db.map_existential ~secvars:Id.Pred.full hdc cl
- else Hint_db.map_auto ~secvars:Id.Pred.full hdc cl
+ let hdc = decompose_app_bound sigma cl in
+ if occur_existential sigma cl then
+ Hint_db.map_existential sigma ~secvars:Id.Pred.full hdc cl
+ else Hint_db.map_auto sigma ~secvars:Id.Pred.full hdc cl
with Bound -> Hint_db.map_none ~secvars:Id.Pred.full
in
let fn db = List.map (fun x -> 0, x) (fn db) in
@@ -1455,7 +1479,7 @@ let pr_applicable_hint () =
match glss.Evd.it with
| [] -> CErrors.error "No focused goal."
| g::_ ->
- pr_hint_term (Goal.V82.concl glss.Evd.sigma g)
+ pr_hint_term glss.Evd.sigma (Goal.V82.concl glss.Evd.sigma g)
let pp_hint_mode = function
| ModeInput -> str"+"
diff --git a/tactics/hints.mli b/tactics/hints.mli
index 05d41adfe..467fd46d5 100644
--- a/tactics/hints.mli
+++ b/tactics/hints.mli
@@ -10,6 +10,7 @@ open Pp
open Util
open Names
open Term
+open EConstr
open Environ
open Globnames
open Decl_kinds
@@ -24,11 +25,11 @@ open Vernacexpr
exception Bound
-val decompose_app_bound : constr -> global_reference * constr array
+val decompose_app_bound : evar_map -> constr -> global_reference * constr array
type debug = Debug | Info | Off
-val secvars_of_hyps : Context.Named.t -> Id.Pred.t
+val secvars_of_hyps : ('c, 't) Context.Named.pt -> Id.Pred.t
val empty_hint_info : 'a hint_info_gen
@@ -110,16 +111,16 @@ module Hint_db :
(** All hints associated to the reference, respecting modes if evars appear in the
arguments, _not_ using the discrimination net. *)
- val map_existential : secvars:Id.Pred.t ->
+ val map_existential : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
(** All hints associated to the reference, respecting modes if evars appear in the
arguments and using the discrimination net. *)
- val map_eauto : secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list
+ val map_eauto : evar_map -> secvars:Id.Pred.t -> (global_reference * constr array) -> constr -> t -> full_hint list
(** All hints associated to the reference, respecting modes if evars appear in the
arguments. *)
- val map_auto : secvars:Id.Pred.t ->
+ val map_auto : evar_map -> secvars:Id.Pred.t ->
(global_reference * constr array) -> constr -> t -> full_hint list
val add_one : env -> evar_map -> hint_entry -> t -> t
@@ -182,7 +183,7 @@ val add_hints : locality_flag -> hint_db_name list -> hints_entry -> unit
val prepare_hint : bool (* Check no remaining evars *) ->
(bool * bool) (* polymorphic or monomorphic, local or global *) ->
- env -> evar_map -> open_constr -> hint_term
+ env -> evar_map -> evar_map * constr -> hint_term
(** [make_exact_entry info (c, ctyp, ctx)].
[c] is the term given as an exact proof to solve the goal;
@@ -231,7 +232,7 @@ val make_resolves :
If the hyp cannot be used as a Hint, the empty list is returned. *)
val make_resolve_hyp :
- env -> evar_map -> Context.Named.Declaration.t -> hint_entry list
+ env -> evar_map -> named_declaration -> hint_entry list
(** [make_extern pri pattern tactic_expr] *)
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index 27af7200b..851554b83 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -12,6 +12,7 @@ open Util
open Names
open Term
open Termops
+open EConstr
open Inductiveops
open Constr_matching
open Coqlib
@@ -31,9 +32,9 @@ module RelDecl = Context.Rel.Declaration
-- Eduardo (6/8/97). *)
-type 'a matching_function = constr -> 'a option
+type 'a matching_function = Evd.evar_map -> EConstr.constr -> 'a option
-type testing_function = constr -> bool
+type testing_function = Evd.evar_map -> EConstr.constr -> bool
let mkmeta n = Nameops.make_ident "X" (Some n)
let meta1 = mkmeta 1
@@ -43,11 +44,11 @@ let meta4 = mkmeta 4
let op2bool = function Some _ -> true | None -> false
-let match_with_non_recursive_type t =
- match kind_of_term t with
+let match_with_non_recursive_type sigma t =
+ match EConstr.kind sigma t with
| App _ ->
- let (hdapp,args) = decompose_app t in
- (match kind_of_term hdapp with
+ let (hdapp,args) = decompose_app sigma t in
+ (match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
if (Global.lookup_mind (fst ind)).mind_finite == Decl_kinds.CoFinite then
Some (hdapp,args)
@@ -56,21 +57,21 @@ let match_with_non_recursive_type t =
| _ -> None)
| _ -> None
-let is_non_recursive_type t = op2bool (match_with_non_recursive_type t)
+let is_non_recursive_type sigma t = op2bool (match_with_non_recursive_type sigma t)
(* Test dependencies *)
(* NB: we consider also the let-in case in the following function,
since they may appear in types of inductive constructors (see #2629) *)
-let rec has_nodep_prod_after n c =
- match kind_of_term c with
+let rec has_nodep_prod_after n sigma c =
+ match EConstr.kind sigma c with
| Prod (_,_,b) | LetIn (_,_,_,b) ->
- ( n>0 || not (dependent (mkRel 1) b))
- && (has_nodep_prod_after (n-1) b)
+ ( n>0 || Vars.noccurn sigma 1 b)
+ && (has_nodep_prod_after (n-1) sigma b)
| _ -> true
-let has_nodep_prod = has_nodep_prod_after 0
+let has_nodep_prod sigma c = has_nodep_prod_after 0 sigma c
(* A general conjunctive type is a non-recursive with-no-indices inductive
type with only one constructor and no dependencies between argument;
@@ -87,9 +88,11 @@ let is_lax_conjunction = function
| Some false -> true
| _ -> false
-let match_with_one_constructor style onlybinary allow_rec t =
- let (hdapp,args) = decompose_app t in
- let res = match kind_of_term hdapp with
+let prod_assum sigma t = fst (decompose_prod_assum sigma t)
+
+let match_with_one_constructor sigma style onlybinary allow_rec t =
+ let (hdapp,args) = decompose_app sigma t in
+ let res = match EConstr.kind sigma hdapp with
| Ind ind ->
let (mib,mip) = Global.lookup_inductive (fst ind) in
if Int.equal (Array.length mip.mind_consnames) 1
@@ -98,21 +101,21 @@ let match_with_one_constructor style onlybinary allow_rec t =
then
if is_strict_conjunction style (* strict conjunction *) then
let ctx =
- (prod_assum (snd
- (decompose_prod_n_assum mib.mind_nparams mip.mind_nf_lc.(0)))) in
+ (prod_assum sigma (snd
+ (decompose_prod_n_assum sigma mib.mind_nparams (EConstr.of_constr mip.mind_nf_lc.(0))))) in
if
List.for_all
(fun decl -> let c = RelDecl.get_type decl in
is_local_assum decl &&
- isRel c &&
- Int.equal (destRel c) mib.mind_nparams) ctx
+ isRel sigma c &&
+ Int.equal (destRel sigma c) mib.mind_nparams) ctx
then
Some (hdapp,args)
else None
else
- let ctyp = prod_applist mip.mind_nf_lc.(0) args in
- let cargs = List.map RelDecl.get_type (prod_assum ctyp) in
- if not (is_lax_conjunction style) || has_nodep_prod ctyp then
+ let ctyp = Termops.prod_applist sigma (EConstr.of_constr mip.mind_nf_lc.(0)) args in
+ let cargs = List.map RelDecl.get_type (prod_assum sigma ctyp) in
+ if not (is_lax_conjunction style) || has_nodep_prod sigma ctyp then
(* Record or non strict conjunction *)
Some (hdapp,List.rev cargs)
else
@@ -125,28 +128,29 @@ let match_with_one_constructor style onlybinary allow_rec t =
| Some (hdapp, [_; _]) -> res
| _ -> None
-let match_with_conjunction ?(strict=false) ?(onlybinary=false) t =
- match_with_one_constructor (Some strict) onlybinary false t
+let match_with_conjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ match_with_one_constructor sigma (Some strict) onlybinary false t
-let match_with_record t =
- match_with_one_constructor None false false t
+let match_with_record sigma t =
+ match_with_one_constructor sigma None false false t
-let is_conjunction ?(strict=false) ?(onlybinary=false) t =
- op2bool (match_with_conjunction ~strict ~onlybinary t)
+let is_conjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ op2bool (match_with_conjunction sigma ~strict ~onlybinary t)
-let is_record t =
- op2bool (match_with_record t)
+let is_record sigma t =
+ op2bool (match_with_record sigma t)
-let match_with_tuple t =
- let t = match_with_one_constructor None false true t in
+let match_with_tuple sigma t =
+ let t = match_with_one_constructor sigma None false true t in
Option.map (fun (hd,l) ->
- let ind = destInd hd in
+ let ind = destInd sigma hd in
+ let ind = on_snd (fun u -> EInstance.kind sigma u) ind in
let (mib,mip) = Global.lookup_pinductive ind in
let isrec = mis_is_recursive (fst ind,mib,mip) in
(hd,l,isrec)) t
-let is_tuple t =
- op2bool (match_with_tuple t)
+let is_tuple sigma t =
+ op2bool (match_with_tuple sigma t)
(* A general disjunction type is a non-recursive with-no-indices inductive
type with of which all constructors have a single argument;
@@ -154,14 +158,15 @@ let is_tuple t =
"Inductive I A1 ... An := C1 (_:A1) | ... | Cn : (_:An)" *)
let test_strict_disjunction n lc =
+ let open Term in
Array.for_all_i (fun i c ->
match (prod_assum (snd (decompose_prod_n_assum n c))) with
| [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i)
| _ -> false) 0 lc
-let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
- let (hdapp,args) = decompose_app t in
- let res = match kind_of_term hdapp with
+let match_with_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ let res = match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
let car = constructors_nrealargs ind in
let (mib,mip) = Global.lookup_inductive ind in
@@ -176,7 +181,7 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
None
else
let cargs =
- Array.map (fun ar -> pi2 (destProd (prod_applist ar args)))
+ Array.map (fun ar -> pi2 (destProd sigma (prod_applist sigma (EConstr.of_constr ar) args)))
mip.mind_nf_lc in
Some (hdapp,Array.to_list cargs)
else
@@ -187,48 +192,48 @@ let match_with_disjunction ?(strict=false) ?(onlybinary=false) t =
| Some (hdapp,[_; _]) -> res
| _ -> None
-let is_disjunction ?(strict=false) ?(onlybinary=false) t =
- op2bool (match_with_disjunction ~strict ~onlybinary t)
+let is_disjunction ?(strict=false) ?(onlybinary=false) sigma t =
+ op2bool (match_with_disjunction ~strict ~onlybinary sigma t)
(* An empty type is an inductive type, possible with indices, that has no
constructors *)
-let match_with_empty_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_empty_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
let nconstr = Array.length mip.mind_consnames in
if Int.equal nconstr 0 then Some hdapp else None
| _ -> None
-let is_empty_type t = op2bool (match_with_empty_type t)
+let is_empty_type sigma t = op2bool (match_with_empty_type sigma t)
(* This filters inductive types with one constructor with no arguments;
Parameters and indices are allowed *)
-let match_with_unit_or_eq_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_unit_or_eq_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind , _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
let constr_types = mip.mind_nf_lc in
let nconstr = Array.length mip.mind_consnames in
- let zero_args c = Int.equal (nb_prod c) mib.mind_nparams in
+ let zero_args c = Int.equal (nb_prod sigma (EConstr.of_constr c)) mib.mind_nparams in
if Int.equal nconstr 1 && zero_args constr_types.(0) then
Some hdapp
else
None
| _ -> None
-let is_unit_or_eq_type t = op2bool (match_with_unit_or_eq_type t)
+let is_unit_or_eq_type sigma t = op2bool (match_with_unit_or_eq_type sigma t)
(* A unit type is an inductive type with no indices but possibly
(useless) parameters, and that has no arguments in its unique
constructor *)
-let is_unit_type t =
- match match_with_conjunction t with
+let is_unit_type sigma t =
+ match match_with_conjunction sigma t with
| Some (_,[]) -> true
| _ -> false
@@ -276,13 +281,13 @@ let coq_refl_jm_pattern =
open Globnames
-let is_matching x y = is_matching (Global.env ()) Evd.empty x y
-let matches x y = matches (Global.env ()) Evd.empty x y
+let is_matching sigma x y = is_matching (Global.env ()) sigma x y
+let matches sigma x y = matches (Global.env ()) sigma x y
-let match_with_equation t =
- if not (isApp t) then raise NoEquationFound;
- let (hdapp,args) = destApp t in
- match kind_of_term hdapp with
+let match_with_equation sigma t =
+ if not (isApp sigma t) then raise NoEquationFound;
+ let (hdapp,args) = destApp sigma t in
+ match EConstr.kind sigma hdapp with
| Ind (ind,u) ->
if eq_gr (IndRef ind) glob_eq then
Some (build_coq_eq_data()),hdapp,
@@ -298,11 +303,11 @@ let match_with_equation t =
let constr_types = mip.mind_nf_lc in
let nconstr = Array.length mip.mind_consnames in
if Int.equal nconstr 1 then
- if is_matching coq_refl_leibniz1_pattern constr_types.(0) then
+ if is_matching sigma coq_refl_leibniz1_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, MonomorphicLeibnizEq(args.(0),args.(1))
- else if is_matching coq_refl_leibniz2_pattern constr_types.(0) then
+ else if is_matching sigma coq_refl_leibniz2_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, PolymorphicLeibnizEq(args.(0),args.(1),args.(2))
- else if is_matching coq_refl_jm_pattern constr_types.(0) then
+ else if is_matching sigma coq_refl_jm_pattern (EConstr.of_constr constr_types.(0)) then
None, hdapp, HeterogenousEq(args.(0),args.(1),args.(2),args.(3))
else raise NoEquationFound
else raise NoEquationFound
@@ -318,57 +323,57 @@ let is_inductive_equality ind =
let nconstr = Array.length mip.mind_consnames in
Int.equal nconstr 1 && Int.equal (constructor_nrealargs (ind,1)) 0
-let match_with_equality_type t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
+let match_with_equality_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
| Ind (ind,_) when is_inductive_equality ind -> Some (hdapp,args)
| _ -> None
-let is_equality_type t = op2bool (match_with_equality_type t)
+let is_equality_type sigma t = op2bool (match_with_equality_type sigma t)
(* Arrows/Implication/Negation *)
(** X1 -> X2 **)
let coq_arrow_pattern = mkPattern (mkGArrow (mkGPatVar "X1") (mkGPatVar "X2"))
-let match_arrow_pattern t =
- let result = matches coq_arrow_pattern t in
+let match_arrow_pattern sigma t =
+ let result = matches sigma coq_arrow_pattern t in
match Id.Map.bindings result with
| [(m1,arg);(m2,mind)] ->
assert (Id.equal m1 meta1 && Id.equal m2 meta2); (arg, mind)
| _ -> anomaly (Pp.str "Incorrect pattern matching")
-let match_with_imp_term c=
- match kind_of_term c with
- | Prod (_,a,b) when not (dependent (mkRel 1) b) ->Some (a,b)
+let match_with_imp_term sigma c =
+ match EConstr.kind sigma c with
+ | Prod (_,a,b) when Vars.noccurn sigma 1 b -> Some (a,b)
| _ -> None
-let is_imp_term c = op2bool (match_with_imp_term c)
+let is_imp_term sigma c = op2bool (match_with_imp_term sigma c)
-let match_with_nottype t =
+let match_with_nottype sigma t =
try
- let (arg,mind) = match_arrow_pattern t in
- if is_empty_type mind then Some (mind,arg) else None
+ let (arg,mind) = match_arrow_pattern sigma t in
+ if is_empty_type sigma mind then Some (mind,arg) else None
with PatternMatchingFailure -> None
-let is_nottype t = op2bool (match_with_nottype t)
+let is_nottype sigma t = op2bool (match_with_nottype sigma t)
(* Forall *)
-let match_with_forall_term c=
- match kind_of_term c with
+let match_with_forall_term sigma c=
+ match EConstr.kind sigma c with
| Prod (nam,a,b) -> Some (nam,a,b)
| _ -> None
-let is_forall_term c = op2bool (match_with_forall_term c)
+let is_forall_term sigma c = op2bool (match_with_forall_term sigma c)
-let match_with_nodep_ind t =
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_nodep_ind sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
if Array.length (mib.mind_packets)>1 then None else
- let nodep_constr = has_nodep_prod_after mib.mind_nparams in
+ let nodep_constr c = has_nodep_prod_after mib.mind_nparams sigma (EConstr.of_constr c) in
if Array.for_all nodep_constr mip.mind_nf_lc then
let params=
if Int.equal mip.mind_nrealargs 0 then args else
@@ -378,24 +383,24 @@ let match_with_nodep_ind t =
None
| _ -> None
-let is_nodep_ind t=op2bool (match_with_nodep_ind t)
+let is_nodep_ind sigma t = op2bool (match_with_nodep_ind sigma t)
-let match_with_sigma_type t=
- let (hdapp,args) = decompose_app t in
- match (kind_of_term hdapp) with
- | Ind ind ->
- let (mib,mip) = Global.lookup_pinductive ind in
+let match_with_sigma_type sigma t =
+ let (hdapp,args) = decompose_app sigma t in
+ match EConstr.kind sigma hdapp with
+ | Ind (ind, _) ->
+ let (mib,mip) = Global.lookup_inductive ind in
if Int.equal (Array.length (mib.mind_packets)) 1 &&
(Int.equal mip.mind_nrealargs 0) &&
(Int.equal (Array.length mip.mind_consnames)1) &&
- has_nodep_prod_after (mib.mind_nparams+1) mip.mind_nf_lc.(0) then
+ has_nodep_prod_after (mib.mind_nparams+1) sigma (EConstr.of_constr mip.mind_nf_lc.(0)) then
(*allowing only 1 existential*)
Some (hdapp,args)
else
None
| _ -> None
-let is_sigma_type t=op2bool (match_with_sigma_type t)
+let is_sigma_type sigma t = op2bool (match_with_sigma_type sigma t)
(***** Destructing patterns bound to some theory *)
@@ -408,17 +413,17 @@ let rec first_match matcher = function
(*** Equality *)
-let match_eq eqn (ref, hetero) =
+let match_eq sigma eqn (ref, hetero) =
let ref =
try Lazy.force ref
with e when CErrors.noncritical e -> raise PatternMatchingFailure
in
- match kind_of_term eqn with
+ match EConstr.kind sigma eqn with
| App (c, [|t; x; y|]) ->
- if not hetero && is_global ref c then PolymorphicLeibnizEq (t, x, y)
+ if not hetero && Termops.is_global sigma ref c then PolymorphicLeibnizEq (t, x, y)
else raise PatternMatchingFailure
| App (c, [|t; x; t'; x'|]) ->
- if hetero && is_global ref c then HeterogenousEq (t, x, t', x')
+ if hetero && Termops.is_global sigma ref c then HeterogenousEq (t, x, t', x')
else raise PatternMatchingFailure
| _ -> raise PatternMatchingFailure
@@ -430,9 +435,9 @@ let equalities =
(coq_jmeq_ref, true), check_jmeq_loaded, build_coq_jmeq_data;
(coq_identity_ref, false), no_check, build_coq_identity_data]
-let find_eq_data eqn = (* fails with PatternMatchingFailure *)
- let d,k = first_match (match_eq eqn) equalities in
- let hd,u = destInd (fst (destApp eqn)) in
+let find_eq_data sigma eqn = (* fails with PatternMatchingFailure *)
+ let d,k = first_match (match_eq sigma eqn) equalities in
+ let hd,u = destInd sigma (fst (destApp sigma eqn)) in
d,u,k
let extract_eq_args gl = function
@@ -444,13 +449,13 @@ let extract_eq_args gl = function
else raise PatternMatchingFailure
let find_eq_data_decompose gl eqn =
- let (lbeq,u,eq_args) = find_eq_data eqn in
+ let (lbeq,u,eq_args) = find_eq_data (project gl) eqn in
(lbeq,u,extract_eq_args gl eq_args)
let find_this_eq_data_decompose gl eqn =
let (lbeq,u,eq_args) =
try (*first_match (match_eq eqn) inversible_equalities*)
- find_eq_data eqn
+ find_eq_data (project gl) eqn
with PatternMatchingFailure ->
user_err (str "No primitive equality found.") in
let eq_args =
@@ -477,12 +482,12 @@ let dest_nf_eq gls eqn =
(*** Sigma-types *)
-let match_sigma ex =
- match kind_of_term ex with
- | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_exist_ref) f ->
- build_sigma (), (snd (destConstruct f), a, p, car, cdr)
- | App (f, [| a; p; car; cdr |]) when is_global (Lazy.force coq_existT_ref) f ->
- build_sigma_type (), (snd (destConstruct f), a, p, car, cdr)
+let match_sigma sigma ex =
+ match EConstr.kind sigma ex with
+ | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_exist_ref) f ->
+ build_sigma (), (snd (destConstruct sigma f), a, p, car, cdr)
+ | App (f, [| a; p; car; cdr |]) when Termops.is_global sigma (Lazy.force coq_existT_ref) f ->
+ build_sigma_type (), (snd (destConstruct sigma f), a, p, car, cdr)
| _ -> raise PatternMatchingFailure
let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *)
@@ -492,12 +497,12 @@ let find_sigma_data_decompose ex = (* fails with PatternMatchingFailure *)
let coq_sig_pattern =
lazy (mkPattern (mkGAppRef coq_sig_ref [mkGPatVar "X1"; mkGPatVar "X2"]))
-let match_sigma t =
- match Id.Map.bindings (matches (Lazy.force coq_sig_pattern) t) with
+let match_sigma sigma t =
+ match Id.Map.bindings (matches sigma (Lazy.force coq_sig_pattern) t) with
| [(_,a); (_,p)] -> (a,p)
| _ -> anomaly (Pp.str "Unexpected pattern")
-let is_matching_sigma t = is_matching (Lazy.force coq_sig_pattern) t
+let is_matching_sigma sigma t = is_matching sigma (Lazy.force coq_sig_pattern) t
(*** Decidable equalities *)
@@ -529,26 +534,26 @@ let coq_eqdec_rev_pattern = coq_eqdec ~sum:coq_or_ref ~rev:true
let op_or = coq_or_ref
let op_sum = coq_sumbool_ref
-let match_eqdec t =
+let match_eqdec sigma t =
let eqonleft,op,subst =
- try true,op_sum,matches (Lazy.force coq_eqdec_inf_pattern) t
+ try true,op_sum,matches sigma (Lazy.force coq_eqdec_inf_pattern) t
with PatternMatchingFailure ->
- try false,op_sum,matches (Lazy.force coq_eqdec_inf_rev_pattern) t
+ try false,op_sum,matches sigma (Lazy.force coq_eqdec_inf_rev_pattern) t
with PatternMatchingFailure ->
- try true,op_or,matches (Lazy.force coq_eqdec_pattern) t
+ try true,op_or,matches sigma (Lazy.force coq_eqdec_pattern) t
with PatternMatchingFailure ->
- false,op_or,matches (Lazy.force coq_eqdec_rev_pattern) t in
+ false,op_or,matches sigma (Lazy.force coq_eqdec_rev_pattern) t in
match Id.Map.bindings subst with
| [(_,typ);(_,c1);(_,c2)] ->
- eqonleft, Universes.constr_of_global (Lazy.force op), c1, c2, typ
+ eqonleft, EConstr.of_constr (Universes.constr_of_global (Lazy.force op)), c1, c2, typ
| _ -> anomaly (Pp.str "Unexpected pattern")
(* Patterns "~ ?" and "? -> False" *)
let coq_not_pattern = lazy (mkPattern (mkGAppRef coq_not_ref [mkGHole]))
let coq_imp_False_pattern = lazy (mkPattern (mkGArrow mkGHole (mkGRef coq_False_ref)))
-let is_matching_not t = is_matching (Lazy.force coq_not_pattern) t
-let is_matching_imp_False t = is_matching (Lazy.force coq_imp_False_pattern) t
+let is_matching_not sigma t = is_matching sigma (Lazy.force coq_not_pattern) t
+let is_matching_imp_False sigma t = is_matching sigma (Lazy.force coq_imp_False_pattern) t
(* Remark: patterns that have references to the standard library must
be evaluated lazily (i.e. at the time they are used, not a the time
diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli
index 7cc41f1b9..dd09c3a4d 100644
--- a/tactics/hipattern.mli
+++ b/tactics/hipattern.mli
@@ -8,6 +8,8 @@
open Names
open Term
+open Evd
+open EConstr
open Coqlib
(** High-order patterns *)
@@ -40,8 +42,8 @@ open Coqlib
also work on ad-hoc disjunctions introduced by the user.
(Eduardo, 6/8/97). *)
-type 'a matching_function = constr -> 'a option
-type testing_function = constr -> bool
+type 'a matching_function = evar_map -> constr -> 'a option
+type testing_function = evar_map -> constr -> bool
val match_with_non_recursive_type : (constr * constr list) matching_function
val is_non_recursive_type : testing_function
@@ -113,39 +115,39 @@ type equation_kind =
exception NoEquationFound
val match_with_equation:
- constr -> coq_eq_data option * constr * equation_kind
+ evar_map -> constr -> coq_eq_data option * constr * equation_kind
(***** Destructing patterns bound to some theory *)
(** Match terms [eq A t u], [identity A t u] or [JMeq A t A u]
Returns associated lemmas and [A,t,u] or fails PatternMatchingFailure *)
-val find_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr ->
- coq_eq_data * Univ.universe_instance * (types * constr * constr)
+val find_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr ->
+ coq_eq_data * EInstance.t * (types * constr * constr)
(** Idem but fails with an error message instead of PatternMatchingFailure *)
-val find_this_eq_data_decompose : ([ `NF ], 'r) Proofview.Goal.t -> constr ->
- coq_eq_data * Univ.universe_instance * (types * constr * constr)
+val find_this_eq_data_decompose : ('a, 'r) Proofview.Goal.t -> constr ->
+ coq_eq_data * EInstance.t * (types * constr * constr)
(** A variant that returns more informative structure on the equality found *)
-val find_eq_data : constr -> coq_eq_data * Univ.universe_instance * equation_kind
+val find_eq_data : evar_map -> constr -> coq_eq_data * EInstance.t * equation_kind
(** Match a term of the form [(existT A P t p)]
Returns associated lemmas and [A,P,t,p] *)
-val find_sigma_data_decompose : constr ->
- coq_sigma_data * (Univ.universe_instance * constr * constr * constr * constr)
+val find_sigma_data_decompose : evar_map -> constr ->
+ coq_sigma_data * (EInstance.t * constr * constr * constr * constr)
(** Match a term of the form [{x:A|P}], returns [A] and [P] *)
-val match_sigma : constr -> constr * constr
+val match_sigma : evar_map -> constr -> constr * constr
-val is_matching_sigma : constr -> bool
+val is_matching_sigma : evar_map -> constr -> bool
(** Match a decidable equality judgement (e.g [{t=u:>T}+{~t=u}]), returns
[t,u,T] and a boolean telling if equality is on the left side *)
-val match_eqdec : constr -> bool * constr * constr * constr * constr
+val match_eqdec : evar_map -> constr -> bool * constr * constr * constr * constr
(** Match an equality up to conversion; returns [(eq,t1,t2)] in normal form *)
-val dest_nf_eq : ([ `NF ], 'r) Proofview.Goal.t -> constr -> (constr * constr * constr)
+val dest_nf_eq : ('a, 'r) Proofview.Goal.t -> constr -> (constr * constr * constr)
(** Match a negation *)
-val is_matching_not : constr -> bool
-val is_matching_imp_False : constr -> bool
+val is_matching_not : evar_map -> constr -> bool
+val is_matching_imp_False : evar_map -> constr -> bool
diff --git a/tactics/inv.ml b/tactics/inv.ml
index e7d8249e4..904a17417 100644
--- a/tactics/inv.ml
+++ b/tactics/inv.ml
@@ -12,10 +12,11 @@ open Util
open Names
open Nameops
open Term
-open Vars
open Termops
-open Namegen
open Environ
+open EConstr
+open Vars
+open Namegen
open Inductiveops
open Printer
open Retyping
@@ -32,8 +33,9 @@ module NamedDecl = Context.Named.Declaration
let var_occurs_in_pf gl id =
let env = Proofview.Goal.env gl in
- occur_var env id (Proofview.Goal.concl gl) ||
- List.exists (occur_var_in_decl env id) (Proofview.Goal.hyps gl)
+ let sigma = project gl in
+ occur_var env sigma id (Proofview.Goal.concl gl) ||
+ List.exists (occur_var_in_decl env sigma id) (Proofview.Goal.hyps gl)
(* [make_inv_predicate (ity,args) C]
@@ -73,9 +75,10 @@ let make_inv_predicate env evd indf realargs id status concl =
| NoDep ->
(* We push the arity and leave concl unchanged *)
let hyps_arity,_ = get_arity env indf in
+ let hyps_arity = List.map (fun d -> map_rel_decl EConstr.of_constr d) hyps_arity in
(hyps_arity,concl)
| Dep dflt_concl ->
- if not (occur_var env id concl) then
+ if not (occur_var env !evd id concl) then
user_err ~hdr:"make_inv_predicate"
(str "Current goal does not depend on " ++ pr_id id ++ str".");
(* We abstract the conclusion of goal with respect to
@@ -87,11 +90,11 @@ let make_inv_predicate env evd indf realargs id status concl =
| None ->
let sort = get_sort_family_of env !evd concl in
let sort = Evarutil.evd_comb1 (Evd.fresh_sort_in_family env) evd sort in
- let p = make_arity env true indf sort in
+ let p = make_arity env !evd true indf sort in
let evd',(p,ptyp) = Unification.abstract_list_all env
!evd p concl (realargs@[mkVar id])
in evd := evd'; p in
- let hyps,bodypred = decompose_lam_n_assum (nrealargs+1) pred in
+ let hyps,bodypred = decompose_lam_n_assum !evd (nrealargs+1) pred in
(* We lift to make room for the equations *)
(hyps,lift nrealargs bodypred)
in
@@ -109,7 +112,7 @@ let make_inv_predicate env evd indf realargs id status concl =
let ai = lift nhyps ai in
let (xi, ti) = compute_eqn env' !evd nhyps n ai in
let (lhs,eqnty,rhs) =
- if closed0 ti then
+ if closed0 !evd ti then
(xi,ti,ai)
else
let sigma, res = make_iterated_tuple env' !evd ai (xi,ti) in
@@ -117,17 +120,19 @@ let make_inv_predicate env evd indf realargs id status concl =
in
let eq_term = eqdata.Coqlib.eq in
let eq = Evarutil.evd_comb1 (Evd.fresh_global env) evd eq_term in
+ let eq = EConstr.of_constr eq in
let eqn = applist (eq,[eqnty;lhs;rhs]) in
let eqns = (Anonymous, lift n eqn) :: eqns in
let refl_term = eqdata.Coqlib.refl in
let refl_term = Evarutil.evd_comb1 (Evd.fresh_global env) evd refl_term in
+ let refl_term = EConstr.of_constr refl_term in
let refl = mkApp (refl_term, [|eqnty; rhs|]) in
let _ = Evarutil.evd_comb1 (Typing.type_of env) evd refl in
let args = refl :: args in
build_concl eqns args (succ n) restlist
in
let (newconcl, args) = build_concl [] [] 0 realargs in
- let predicate = it_mkLambda_or_LetIn_name env newconcl hyps in
+ let predicate = it_mkLambda_or_LetIn newconcl (name_context env !evd hyps) in
let _ = Evarutil.evd_comb1 (Typing.type_of env) evd predicate in
(* OK - this predicate should now be usable by res_elimination_then to
do elimination on the conclusion. *)
@@ -183,7 +188,7 @@ let dependent_hyps env id idlist gl =
| d::l ->
(* Update the type of id1: it may have been subject to rewriting *)
let d = pf_get_hyp (NamedDecl.get_id d) gl in
- if occur_var_in_decl env id d
+ if occur_var_in_decl env (project gl) id d
then d :: dep_rec l
else dep_rec l
in
@@ -268,7 +273,7 @@ Nota: with Inversion_clear, only four useless hypotheses
let generalizeRewriteIntros as_mode tac depids id =
Proofview.tclENV >>= fun env ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let dids = dependent_hyps env id depids gl in
let reintros = if as_mode then intros_replacing else intros_possibly_replacing in
(tclTHENLIST
@@ -283,7 +288,7 @@ let error_too_many_names pats =
tclZEROMSG ~loc (
str "Unexpected " ++
str (String.plural (List.length pats) "introduction pattern") ++
- str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (fst (run_delayed env Evd.empty c)))) pats ++
+ str ": " ++ pr_enum (Miscprint.pr_intro_pattern (fun c -> Printer.pr_constr (EConstr.Unsafe.to_constr (fst (run_delayed env Evd.empty c))))) pats ++
str ".")
let get_names (allow_conj,issimple) (loc, pat as x) = match pat with
@@ -333,15 +338,16 @@ let remember_first_eq id x = if !x == MoveLast then x := MoveAfter id
let projectAndApply as_mode thin avoid id eqname names depids =
let subst_hyp l2r id =
- tclTHEN (tclTRY(rewriteInConcl l2r (mkVar id)))
+ tclTHEN (tclTRY(rewriteInConcl l2r (EConstr.mkVar id)))
(if thin then clear [id] else (remember_first_eq id eqname; tclIDTAC))
in
let substHypIfVariable tac id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = project gl in
(** We only look at the type of hypothesis "id" *)
let hyp = pf_nf_evar gl (pf_get_hyp_typ id (Proofview.Goal.assume gl)) in
let (t,t1,t2) = Hipattern.dest_nf_eq gl hyp in
- match (kind_of_term t1, kind_of_term t2) with
+ match (EConstr.kind sigma t1, EConstr.kind sigma t2) with
| Var id1, _ -> generalizeRewriteIntros as_mode (subst_hyp true id) depids id1
| _, Var id2 -> generalizeRewriteIntros as_mode (subst_hyp false id) depids id2
| _ -> tac id
@@ -368,11 +374,11 @@ let projectAndApply as_mode thin avoid id eqname names depids =
(* and apply a trailer which again try to substitute *)
(fun id ->
dEqThen false (deq_trailer id)
- (Some (None,ElimOnConstr (mkVar id,NoBindings))))
+ (Some (None,ElimOnConstr (EConstr.mkVar id,NoBindings))))
id
let nLastDecls i tac =
- Proofview.Goal.nf_enter { enter = begin fun gl -> tac (nLastDecls gl i) end }
+ Proofview.Goal.enter { enter = begin fun gl -> tac (nLastDecls gl i) end }
(* Introduction of the equations on arguments
othin: discriminates Simple Inversion, Inversion and Inversion_clear
@@ -380,7 +386,7 @@ let nLastDecls i tac =
Some thin: the equations are rewritten, and cleared if thin is true *)
let rewrite_equations as_mode othin neqns names ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let (depids,nodepids) = split_dep_and_nodep ba.Tacticals.assums gl in
let first_eq = ref MoveLast in
let avoid = if as_mode then List.map NamedDecl.get_id nodepids else [] in
@@ -430,7 +436,7 @@ let rewrite_equations_tac as_mode othin id neqns names ba =
tac
let raw_inversion inv_kind id status names =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let sigma = Sigma.to_evar_map sigma in
let env = Proofview.Goal.env gl in
@@ -448,11 +454,11 @@ let raw_inversion inv_kind id status names =
make_inv_predicate env evdref indf realargs id status concl in
let sigma = !evdref in
let (cut_concl,case_tac) =
- if status != NoDep && (dependent c concl) then
- Reduction.beta_appvect elim_predicate (Array.of_list (realargs@[c])),
+ if status != NoDep && (dependent sigma c concl) then
+ Reductionops.beta_applist sigma (elim_predicate, realargs@[c]),
case_then_using
else
- Reduction.beta_appvect elim_predicate (Array.of_list realargs),
+ Reductionops.beta_applist sigma (elim_predicate, realargs),
case_nodep_then_using
in
let refined id =
@@ -467,7 +473,7 @@ let raw_inversion inv_kind id status names =
[case_tac names
(introCaseAssumsThen false (* ApplyOn not supported by inversion *)
(rewrite_equations_tac as_mode inv_kind id neqns))
- (Some elim_predicate) ind (c, t);
+ (Some elim_predicate) ind (c,t);
onLastHypId (fun id -> tclTHEN (refined id) reflexivity)])
in
Sigma.Unsafe.of_pair (tac, sigma)
@@ -511,15 +517,17 @@ let dinv_clear_tac id = dinv FullInversionClear None None (NamedHyp id)
* back to their places in the hyp-list. *)
let invIn k names ids id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let hyps = List.map (fun id -> pf_get_hyp id gl) ids in
let concl = Proofview.Goal.concl gl in
- let nb_prod_init = nb_prod concl in
+ let sigma = project gl in
+ let nb_prod_init = nb_prod sigma concl in
let intros_replace_ids =
Proofview.Goal.enter { enter = begin fun gl ->
- let concl = pf_nf_concl gl in
+ let concl = pf_concl gl in
+ let sigma = project gl in
let nb_of_new_hyp =
- nb_prod concl - (List.length hyps + nb_prod_init)
+ nb_prod sigma concl - (List.length hyps + nb_prod_init)
in
if nb_of_new_hyp < 1 then
intros_replacing ids
diff --git a/tactics/inv.mli b/tactics/inv.mli
index df629e7c9..446a62f6d 100644
--- a/tactics/inv.mli
+++ b/tactics/inv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Misctypes
open Tactypes
diff --git a/tactics/leminv.ml b/tactics/leminv.ml
index 10fc5076c..daa962f1d 100644
--- a/tactics/leminv.ml
+++ b/tactics/leminv.ml
@@ -11,15 +11,16 @@ open CErrors
open Util
open Names
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Namegen
open Evd
open Printer
open Reductionops
open Entries
open Inductiveops
-open Environ
open Tacmach.New
open Clenv
open Declare
@@ -33,7 +34,7 @@ module NamedDecl = Context.Named.Declaration
let no_inductive_inconstr env sigma constr =
(str "Cannot recognize an inductive predicate in " ++
- pr_lconstr_env env sigma constr ++
+ pr_leconstr_env env sigma constr ++
str "." ++ spc () ++ str "If there is one, may be the structure of the arity" ++
spc () ++ str "or of the type of constructors" ++ spc () ++
str "is hidden by constant definitions.")
@@ -116,13 +117,13 @@ let max_prefix_sign lid sign =
| id::l -> snd (max_rec (id, sign_prefix id sign) l)
*)
let rec add_prods_sign env sigma t =
- match kind_of_term (whd_all env sigma t) with
+ match EConstr.kind sigma (whd_all env sigma t) with
| Prod (na,c1,b) ->
- let id = id_of_name_using_hdchar env t na in
+ let id = id_of_name_using_hdchar env sigma t na in
let b'= subst1 (mkVar id) b in
add_prods_sign (push_named (LocalAssum (id,c1)) env) sigma b'
| LetIn (na,c1,t1,b) ->
- let id = id_of_name_using_hdchar env t na in
+ let id = id_of_name_using_hdchar env sigma t na in
let b'= subst1 (mkVar id) b in
add_prods_sign (push_named (LocalDef (id,c1,t1)) env) sigma b'
| _ -> (env,t)
@@ -146,7 +147,7 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
let p = next_ident_away (Id.of_string "P") allvars in
let pty,goal =
if dep_option then
- let pty = make_arity env true indf sort in
+ let pty = make_arity env sigma true indf sort in
let goal =
mkProd
(Anonymous, mkAppliedInd ind, applist(mkVar p,realargs@[mkRel 1]))
@@ -154,10 +155,11 @@ let compute_first_inversion_scheme env sigma ind sort dep_option =
pty,goal
else
let i = mkAppliedInd ind in
- let ivars = global_vars env i in
+ let ivars = global_vars env sigma i in
let revargs,ownsign =
fold_named_context
(fun env d (revargs,hyps) ->
+ let d = map_named_decl EConstr.of_constr d in
let id = NamedDecl.get_id d in
if Id.List.mem id ivars then
((mkVar id)::revargs, Context.Named.add d hyps)
@@ -192,7 +194,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
in
assert
(List.subset
- (global_vars env invGoal)
+ (global_vars env sigma invGoal)
(ids_of_named_context (named_context invEnv)));
(*
user_err ~hdr:"lemma_inversion"
@@ -208,6 +210,7 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let ownSign = ref begin
fold_named_context
(fun env d sign ->
+ let d = map_named_decl EConstr.of_constr d in
if mem_named_context_val (NamedDecl.get_id d) global_named_context then sign
else Context.Named.add d sign)
invEnv ~init:Context.Named.empty
@@ -216,18 +219,19 @@ let inversion_scheme env sigma t sort dep_option inv_op =
let { sigma=sigma } = Proof.V82.subgoals pf in
let sigma = Evd.nf_constraints sigma in
let rec fill_holes c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Evar (e,args) ->
let h = next_ident_away (Id.of_string "H") !avoid in
let ty,inst = Evarutil.generalize_evar_over_rels sigma (e,args) in
avoid := h::!avoid;
ownSign := Context.Named.add (LocalAssum (h,ty)) !ownSign;
applist (mkVar h, inst)
- | _ -> Constr.map fill_holes c
+ | _ -> EConstr.map sigma fill_holes c
in
let c = fill_holes pfterm in
(* warning: side-effect on ownSign *)
let invProof = it_mkNamedLambda_or_LetIn c !ownSign in
+ let invProof = EConstr.Unsafe.to_constr invProof in
let p = Evarutil.nf_evars_universes sigma invProof in
p, Evd.universe_context sigma
@@ -258,26 +262,28 @@ let add_inversion_lemma_exn na com comsort bool tac =
let lemInv id c gls =
try
- let clause = mk_clenv_type_of gls c in
- let clause = clenv_constrain_last_binding (mkVar id) clause in
+ let open Tacmach in
+ let clause = mk_clenv_from_env (pf_env gls) (project gls) None (c, pf_unsafe_type_of gls c) in
+ let clause = clenv_constrain_last_binding (EConstr.mkVar id) clause in
Proofview.V82.of_tactic (Clenvtac.res_pf clause ~flags:(Unification.elim_flags ()) ~with_evars:false) gls
with
| NoSuchBinding ->
user_err
- (hov 0 (pr_constr c ++ spc () ++ str "does not refer to an inversion lemma."))
+ (hov 0 (pr_econstr_env (Refiner.pf_env gls) (Refiner.project gls) c ++ spc () ++ str "does not refer to an inversion lemma."))
| UserError (a,b) ->
user_err ~hdr:"LemInv"
(str "Cannot refine current goal with the lemma " ++
- pr_lconstr_env (Refiner.pf_env gls) (Refiner.project gls) c)
+ pr_leconstr_env (Refiner.pf_env gls) (Refiner.project gls) c)
let lemInv_gen id c = try_intros_until (fun id -> Proofview.V82.tactic (lemInv id c)) id
let lemInvIn id c ids =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let hyps = List.map (fun id -> pf_get_hyp id gl) ids in
let intros_replace_ids =
let concl = Proofview.Goal.concl gl in
- let nb_of_new_hyp = nb_prod concl - List.length ids in
+ let sigma = project gl in
+ let nb_of_new_hyp = nb_prod sigma concl - List.length ids in
if nb_of_new_hyp < 1 then
intros_replacing ids
else
diff --git a/tactics/leminv.mli b/tactics/leminv.mli
index c6ed9606f..26d4ac994 100644
--- a/tactics/leminv.mli
+++ b/tactics/leminv.mli
@@ -8,6 +8,7 @@
open Names
open Term
+open EConstr
open Constrexpr
open Misctypes
diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml
index c5562b326..90b7d6581 100644
--- a/tactics/tacticals.ml
+++ b/tactics/tacticals.ml
@@ -11,6 +11,7 @@ open CErrors
open Util
open Names
open Term
+open EConstr
open Termops
open Declarations
open Tacmach
@@ -83,7 +84,7 @@ let nLastDecls n gl =
try List.firstn n (pf_hyps gl)
with Failure _ -> error "Not enough hypotheses in the goal."
-let nLastHypsId n gl = List.map NamedDecl.get_id (nLastDecls n gl)
+let nLastHypsId n gl = List.map (NamedDecl.get_id) (nLastDecls n gl)
let nLastHyps n gl = List.map mkVar (nLastHypsId n gl)
let onNthDecl m tac gl = tac (nthDecl m gl) gl
@@ -157,7 +158,7 @@ type branch_args = {
type branch_assumptions = {
ba : branch_args; (* the branch args *)
- assums : Context.Named.t} (* the list of assumptions introduced *)
+ assums : named_context} (* the list of assumptions introduced *)
open Misctypes
@@ -226,6 +227,7 @@ let compute_induction_names = compute_induction_names_gen true
(* Compute the let-in signature of case analysis or standard induction scheme *)
let compute_constructor_signatures isrec ((_,k as ity),u) =
+ let open Term in
let rec analrec c recargs =
match kind_of_term c, recargs with
| Prod (_,_,c), recarg::rest ->
@@ -263,40 +265,7 @@ let pf_with_evars glsev k gls =
tclTHEN (Refiner.tclEVARS evd) (k a) gls
let pf_constr_of_global gr k =
- pf_with_evars (fun gls -> pf_apply Evd.fresh_global gls gr) k
-
-(* computing the case/elim combinators *)
-
-let gl_make_elim ind gl =
- let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in
- pf_apply Evd.fresh_global gl gr
-
-let gl_make_case_dep ind gl =
- let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in
- let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind true
- (elimination_sort_of_goal gl)
- in
- (Sigma.to_evar_map sigma, r)
-
-let gl_make_case_nodep ind gl =
- let sigma = Sigma.Unsafe.of_evar_map (Tacmach.project gl) in
- let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma ind false
- (elimination_sort_of_goal gl)
- in
- (Sigma.to_evar_map sigma, r)
-
-let make_elim_branch_assumptions ba gl =
- let assums =
- try List.rev (List.firstn ba.nassums (pf_hyps gl))
- with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in
- { ba = ba; assums = assums }
-
-let elim_on_ba tac ba gl = tac (make_elim_branch_assumptions ba gl) gl
-
-let make_case_branch_assumptions = make_elim_branch_assumptions
-
-let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl
-
+ pf_with_evars (fun gls -> on_snd EConstr.of_constr (pf_apply Evd.fresh_global gls gr)) k
(** Tacticals of Ltac defined directly in term of Proofview *)
module New = struct
@@ -544,7 +513,7 @@ module New = struct
Proofview.Unsafe.tclEVARS sigma <*> tac >>= check_evars_if
let tclDELAYEDWITHHOLES check x tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
let Sigma (x, sigma, _) = x.delayed env sigma in
@@ -588,13 +557,13 @@ module New = struct
let onLastHyp = onNthHyp 1
let onNthDecl m tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Proofview.tclUNIT (nthDecl m gl) >>= tac
end }
let onLastDecl = onNthDecl 1
let ifOnHyp pred tac1 tac2 id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let typ = Tacmach.New.pf_get_hyp_typ id gl in
if pred (id,typ) then
tac1 id
@@ -602,7 +571,7 @@ module New = struct
tac2 id
end }
- let onHyps find tac = Proofview.Goal.nf_enter { enter = begin fun gl -> tac (find.enter gl) end }
+ let onHyps find tac = Proofview.Goal.enter { enter = begin fun gl -> tac (find.enter gl) end }
let afterHyp id tac =
Proofview.Goal.enter { enter = begin fun gl ->
@@ -635,25 +604,26 @@ module New = struct
(* c should be of type A1->.. An->B with B an inductive definition *)
let general_elim_then_using mk_elim
isrec allnames tac predicate ind (c, t) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let sigma, elim = Tacmach.New.of_old (mk_elim ind) gl in
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma, elim = (mk_elim ind).enter gl in
+ let ind = on_snd (fun u -> EInstance.kind sigma u) ind in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
- (Proofview.Goal.nf_enter { enter = begin fun gl ->
- let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from gl (c, t)) gl in
+ (Proofview.Goal.enter { enter = begin fun gl ->
+ let indclause = mk_clenv_from gl (c, t) in
(* applying elimination_scheme just a little modified *)
- let elimclause = Tacmach.New.of_old (fun gls -> mk_clenv_from gls (elim,Tacmach.New.pf_unsafe_type_of gl elim)) gl in
+ let elimclause = mk_clenv_from gl (elim,Tacmach.New.pf_unsafe_type_of gl elim) in
let indmv =
- match kind_of_term (last_arg elimclause.templval.Evd.rebus) with
+ match EConstr.kind elimclause.evd (last_arg elimclause.evd elimclause.templval.Evd.rebus) with
| Meta mv -> mv
| _ -> anomaly (str"elimination")
in
let pmv =
- let p, _ = decompose_app elimclause.templtyp.Evd.rebus in
- match kind_of_term p with
+ let p, _ = decompose_app elimclause.evd elimclause.templtyp.Evd.rebus in
+ match EConstr.kind elimclause.evd p with
| Meta p -> p
| _ ->
let name_elim =
- match kind_of_term elim with
+ match EConstr.kind sigma elim with
| Const (kn, _) -> string_of_con kn
| Var id -> string_of_id id
| _ -> "\b"
@@ -670,9 +640,9 @@ module New = struct
| None -> elimclause'
| Some p -> clenv_unify ~flags Reduction.CONV (mkMeta pmv) p elimclause'
in
- let clenv' = Tacmach.New.of_old (clenv_unique_resolver ~flags elimclause') gl in
+ let clenv' = clenv_unique_resolver ~flags elimclause' gl in
let after_tac i =
- let (hd,largs) = decompose_app clenv'.templtyp.Evd.rebus in
+ let (hd,largs) = decompose_app clenv'.evd clenv'.templtyp.Evd.rebus in
let ba = { branchsign = branchsigns.(i);
branchnames = brnames.(i);
nassums = List.length branchsigns.(i);
@@ -689,8 +659,66 @@ module New = struct
(Proofview.tclEXTEND [] tclIDTAC branchtacs)
end }) end }
+ let elimination_sort_of_goal gl =
+ (** Retyping will expand evars anyway. *)
+ let c = Proofview.Goal.concl (Goal.assume gl) in
+ pf_apply Retyping.get_sort_family_of gl c
+
+ let elimination_sort_of_hyp id gl =
+ (** Retyping will expand evars anyway. *)
+ let c = pf_get_hyp_typ id (Goal.assume gl) in
+ pf_apply Retyping.get_sort_family_of gl c
+
+ let elimination_sort_of_clause id gl = match id with
+ | None -> elimination_sort_of_goal gl
+ | Some id -> elimination_sort_of_hyp id gl
+
+ (* computing the case/elim combinators *)
+
+ let gl_make_elim ind = { enter = begin fun gl ->
+ let gr = Indrec.lookup_eliminator (fst ind) (elimination_sort_of_goal gl) in
+ let (sigma, c) = pf_apply Evd.fresh_global gl gr in
+ (sigma, EConstr.of_constr c)
+ end }
+
+ let gl_make_case_dep (ind, u) = { enter = begin fun gl ->
+ let sigma = Sigma.Unsafe.of_evar_map (project gl) in
+ let u = EInstance.kind (project gl) u in
+ let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) true
+ (elimination_sort_of_goal gl)
+ in
+ (Sigma.to_evar_map sigma, EConstr.of_constr r)
+ end }
+
+ let gl_make_case_nodep (ind, u) = { enter = begin fun gl ->
+ let sigma = Sigma.Unsafe.of_evar_map (project gl) in
+ let u = EInstance.kind (project gl) u in
+ let Sigma (r, sigma, _) = Indrec.build_case_analysis_scheme (pf_env gl) sigma (ind, u) false
+ (elimination_sort_of_goal gl)
+ in
+ (Sigma.to_evar_map sigma, EConstr.of_constr r)
+ end }
+
+ let make_elim_branch_assumptions ba hyps =
+ let assums =
+ try List.rev (List.firstn ba.nassums hyps)
+ with Failure _ -> anomaly (Pp.str "make_elim_branch_assumptions") in
+ { ba = ba; assums = assums }
+
+ let elim_on_ba tac ba =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in
+ tac branches
+ end }
+
+ let case_on_ba tac ba =
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let branches = make_elim_branch_assumptions ba (Proofview.Goal.hyps gl) in
+ tac branches
+ end }
+
let elimination_then tac c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let (ind,t) = pf_reduce_to_quantified_ind gl (pf_unsafe_type_of gl c) in
let isrec,mkelim =
match (Global.lookup_mind (fst (fst ind))).mind_record with
@@ -706,37 +734,12 @@ module New = struct
let case_nodep_then_using =
general_elim_then_using gl_make_case_nodep false
- let elim_on_ba tac ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let branches = Tacmach.New.of_old (make_elim_branch_assumptions ba) gl in
- tac branches
- end }
-
- let case_on_ba tac ba =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let branches = Tacmach.New.of_old (make_case_branch_assumptions ba) gl in
- tac branches
- end }
-
- let elimination_sort_of_goal gl =
- (** Retyping will expand evars anyway. *)
- let c = Proofview.Goal.concl (Goal.assume gl) in
- pf_apply Retyping.get_sort_family_of gl c
-
- let elimination_sort_of_hyp id gl =
- (** Retyping will expand evars anyway. *)
- let c = pf_get_hyp_typ id (Goal.assume gl) in
- pf_apply Retyping.get_sort_family_of gl c
-
- let elimination_sort_of_clause id gl = match id with
- | None -> elimination_sort_of_goal gl
- | Some id -> elimination_sort_of_hyp id gl
-
let pf_constr_of_global ref tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let (sigma, c) = Evd.fresh_global env sigma ref in
+ let c = EConstr.of_constr c in
Proofview.Unsafe.tclEVARS sigma <*> (tac c)
end }
diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli
index 7aacc52f3..3b90ec514 100644
--- a/tactics/tacticals.mli
+++ b/tactics/tacticals.mli
@@ -9,6 +9,7 @@
open Pp
open Names
open Term
+open EConstr
open Tacmach
open Proof_type
open Locus
@@ -59,29 +60,29 @@ val tclIFTHENTRYELSEMUST : tactic -> tactic -> tactic
val onNthHypId : int -> (Id.t -> tactic) -> tactic
val onNthHyp : int -> (constr -> tactic) -> tactic
-val onNthDecl : int -> (Context.Named.Declaration.t -> tactic) -> tactic
+val onNthDecl : int -> (named_declaration -> tactic) -> tactic
val onLastHypId : (Id.t -> tactic) -> tactic
val onLastHyp : (constr -> tactic) -> tactic
-val onLastDecl : (Context.Named.Declaration.t -> tactic) -> tactic
+val onLastDecl : (named_declaration -> tactic) -> tactic
val onNLastHypsId : int -> (Id.t list -> tactic) -> tactic
val onNLastHyps : int -> (constr list -> tactic) -> tactic
-val onNLastDecls : int -> (Context.Named.t -> tactic) -> tactic
+val onNLastDecls : int -> (named_context -> tactic) -> tactic
val lastHypId : goal sigma -> Id.t
val lastHyp : goal sigma -> constr
-val lastDecl : goal sigma -> Context.Named.Declaration.t
+val lastDecl : goal sigma -> named_declaration
val nLastHypsId : int -> goal sigma -> Id.t list
val nLastHyps : int -> goal sigma -> constr list
-val nLastDecls : int -> goal sigma -> Context.Named.t
+val nLastDecls : int -> goal sigma -> named_context
-val afterHyp : Id.t -> goal sigma -> Context.Named.t
+val afterHyp : Id.t -> goal sigma -> named_context
val ifOnHyp : (Id.t * types -> bool) ->
(Id.t -> tactic) -> (Id.t -> tactic) ->
Id.t -> tactic
-val onHyps : (goal sigma -> Context.Named.t) ->
- (Context.Named.t -> tactic) -> tactic
+val onHyps : (goal sigma -> named_context) ->
+ (named_context -> tactic) -> tactic
(** {6 Tacticals applying to goal components } *)
@@ -97,7 +98,7 @@ val onClauseLR : (Id.t option -> tactic) -> clause -> tactic
(** {6 Elimination tacticals. } *)
-type branch_args = {
+type branch_args = private {
ity : pinductive; (** the type we were eliminating on *)
largs : constr list; (** its arguments *)
branchnum : int; (** the branch number *)
@@ -107,9 +108,9 @@ type branch_args = {
true=assumption, false=let-in *)
branchnames : intro_patterns}
-type branch_assumptions = {
+type branch_assumptions = private {
ba : branch_args; (** the branch args *)
- assums : Context.Named.t} (** the list of assumptions introduced *)
+ assums : named_context} (** the list of assumptions introduced *)
(** [get_and_check_or_and_pattern loc pats branchsign] returns an appropriate
error message if |pats| <> |branchsign|; extends them if no pattern is given
@@ -123,7 +124,7 @@ val fix_empty_or_and_pattern : int ->
delayed_open_constr or_and_intro_pattern_expr ->
delayed_open_constr or_and_intro_pattern_expr
-val compute_constructor_signatures : rec_flag -> pinductive -> bool list array
+val compute_constructor_signatures : rec_flag -> inductive * 'a -> bool list array
(** Useful for [as intro_pattern] modifier *)
val compute_induction_names :
@@ -136,9 +137,6 @@ val elimination_sort_of_clause : Id.t option -> goal sigma -> sorts_family
val pf_with_evars : (goal sigma -> Evd.evar_map * 'a) -> ('a -> tactic) -> tactic
val pf_constr_of_global : Globnames.global_reference -> (constr -> tactic) -> tactic
-val elim_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
-val case_on_ba : (branch_assumptions -> tactic) -> branch_args -> tactic
-
(** Tacticals defined directly in term of Proofview *)
(** The tacticals in the module [New] are the tactical of Ltac. Their
@@ -229,7 +227,7 @@ module New : sig
val tclTIMEOUT : int -> unit tactic -> unit tactic
val tclTIME : string option -> 'a tactic -> 'a tactic
- val nLastDecls : ([ `NF ], 'r) Proofview.Goal.t -> int -> Context.Named.t
+ val nLastDecls : ('a, 'r) Proofview.Goal.t -> int -> named_context
val ifOnHyp : (identifier * types -> bool) ->
(identifier -> unit Proofview.tactic) -> (identifier -> unit Proofview.tactic) ->
@@ -238,11 +236,11 @@ module New : sig
val onNthHypId : int -> (identifier -> unit tactic) -> unit tactic
val onLastHypId : (identifier -> unit tactic) -> unit tactic
val onLastHyp : (constr -> unit tactic) -> unit tactic
- val onLastDecl : (Context.Named.Declaration.t -> unit tactic) -> unit tactic
+ val onLastDecl : (named_declaration -> unit tactic) -> unit tactic
- val onHyps : ([ `NF ], Context.Named.t) Proofview.Goal.enter ->
- (Context.Named.t -> unit tactic) -> unit tactic
- val afterHyp : Id.t -> (Context.Named.t -> unit tactic) -> unit tactic
+ val onHyps : ([ `LZ ], named_context) Proofview.Goal.enter ->
+ (named_context -> unit tactic) -> unit tactic
+ val afterHyp : Id.t -> (named_context -> unit tactic) -> unit tactic
val tryAllHyps : (identifier -> unit tactic) -> unit tactic
val tryAllHypsAndConcl : (identifier option -> unit tactic) -> unit tactic
@@ -258,11 +256,11 @@ module New : sig
val case_then_using :
or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) ->
- constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic
+ constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic
val case_nodep_then_using :
or_and_intro_pattern option -> (branch_args -> unit Proofview.tactic) ->
- constr option -> pinductive -> Term.constr * Term.types -> unit Proofview.tactic
+ constr option -> inductive * EInstance.t -> constr * types -> unit Proofview.tactic
val elim_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
val case_on_ba : (branch_assumptions -> unit Proofview.tactic) -> branch_args -> unit Proofview.tactic
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 0aab77314..e79258582 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -6,20 +6,23 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+module CVars = Vars
+
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open Environ
+open EConstr
+open Vars
open Find_subterm
open Namegen
open Declarations
open Inductiveops
open Reductionops
-open Environ
open Globnames
open Evd
open Pfedit
@@ -193,7 +196,7 @@ let introduction ?(check=true) id =
(str "Variable " ++ pr_id id ++ str " is already declared.")
in
let open Context.Named.Declaration in
- match kind_of_term (whd_evar sigma concl) with
+ match EConstr.kind sigma concl with
| Prod (_, t, b) -> unsafe_intro env store (LocalAssum (id, t)) b
| LetIn (_, c, t, b) -> unsafe_intro env store (LocalDef (id, c, t)) b
| _ -> raise (RefinerError IntroNeedsProduct)
@@ -205,7 +208,7 @@ let convert_concl ?(check=true) ty k =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- let conclty = Proofview.Goal.raw_concl gl in
+ let conclty = Proofview.Goal.concl gl in
Refine.refine ~unsafe:true { run = begin fun sigma ->
let Sigma ((), sigma, p) =
if check then begin
@@ -225,7 +228,7 @@ let convert_hyp ?(check=true) d =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let ty = Proofview.Goal.raw_concl gl in
+ let ty = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let sign = convert_hyp check (named_context_val env) sigma d in
let env = reset_with_named_context sign env in
@@ -309,24 +312,26 @@ let clear ids = clear_gen error_clear_dependency ids
let clear_for_replacing ids = clear_gen error_replacing_dependency ids
let apply_clear_request clear_flag dft c =
+ Proofview.tclEVARMAP >>= fun sigma ->
let check_isvar c =
- if not (isVar c) then
+ if not (isVar sigma c) then
error "keep/clear modifiers apply only to hypothesis names." in
let doclear = match clear_flag with
- | None -> dft && isVar c
+ | None -> dft && isVar sigma c
| Some true -> check_isvar c; true
| Some false -> false in
- if doclear then clear [destVar c]
+ if doclear then clear [destVar sigma c]
else Tacticals.New.tclIDTAC
(* Moving hypotheses *)
let move_hyp id dest =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
- let ty = Proofview.Goal.raw_concl gl in
+ let sigma = Tacmach.New.project gl in
+ let ty = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let sign = named_context_val env in
- let sign' = move_hyp_in_named_context id dest sign in
+ let sign' = move_hyp_in_named_context sigma id dest sign in
let env = reset_with_named_context sign' env in
Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar env sigma ~principal:true ~store ty
@@ -380,7 +385,7 @@ let rename_hyp repl =
in
let nhyps = List.map map hyps in
let nconcl = subst concl in
- let nctx = Environ.val_of_named_context nhyps in
+ let nctx = val_of_named_context nhyps in
let instance = List.map (NamedDecl.get_id %> mkVar) hyps in
Refine.refine ~unsafe:true { run = begin fun sigma ->
Evarutil.new_evar_instance nctx sigma nconcl ~principal:true ~store instance
@@ -413,7 +418,7 @@ let default_id env sigma decl =
| LocalAssum (name,t) ->
let dft = default_id_of_sort (Retyping.get_sort_of env sigma t) in
id_of_name_with_default dft name
- | LocalDef (name,b,_) -> id_of_name_using_hdchar env b name
+ | LocalDef (name,b,_) -> id_of_name_using_hdchar env sigma b name
(* Non primitive introduction tactics are treated by intro_then_gen
There is possibly renaming, with possibly names to avoid and
@@ -497,7 +502,7 @@ fun env sigma p -> function
let Sigma (rem, sigma, r) = mk_holes env sigma (p +> q) rem in
Sigma (arg :: rem, sigma, r)
-let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) with
+let rec check_mutind env sigma k cl = match EConstr.kind sigma (strip_outer_cast sigma cl) with
| Prod (na, c1, b) ->
if Int.equal k 1 then
try
@@ -510,7 +515,7 @@ let rec check_mutind env sigma k cl = match kind_of_term (strip_outer_cast cl) w
| _ -> error "Not enough products."
(* Refine as a fixpoint *)
-let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let mutual_fix f n rest j = Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
@@ -538,7 +543,7 @@ let mutual_fix f n rest j = Proofview.Goal.nf_enter { enter = begin fun gl ->
let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
let typarray = Array.of_list (List.map pi3 all) in
let bodies = Array.of_list evs in
- let oterm = Term.mkFix ((indxs,0),(funnames,typarray,bodies)) in
+ let oterm = mkFix ((indxs,0),(funnames,typarray,bodies)) in
Sigma (oterm, sigma, p)
end }
end }
@@ -555,7 +560,7 @@ let fix ido n = match ido with
let rec check_is_mutcoind env sigma cl =
let b = whd_all env sigma cl in
- match kind_of_term b with
+ match EConstr.kind sigma b with
| Prod (na, c1, b) ->
let open Context.Rel.Declaration in
check_is_mutcoind (push_rel (LocalAssum (na,c1)) env) sigma b
@@ -566,7 +571,7 @@ let rec check_is_mutcoind env sigma cl =
error "All methods must construct elements in coinductive types."
(* Refine as a cofixpoint *)
-let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let mutual_cofix f others j = Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl gl in
@@ -589,7 +594,7 @@ let mutual_cofix f others j = Proofview.Goal.nf_enter { enter = begin fun gl ->
let funnames = Array.of_list (List.map (fun i -> Name i) ids) in
let typarray = Array.of_list types in
let bodies = Array.of_list evs in
- let oterm = Term.mkCoFix (0, (funnames, typarray, bodies)) in
+ let oterm = mkCoFix (0, (funnames, typarray, bodies)) in
Sigma (oterm, sigma, p)
end }
end }
@@ -612,7 +617,7 @@ type tactic_reduction = env -> evar_map -> constr -> constr
let pf_reduce_decl redfun where decl gl =
let open Context.Named.Declaration in
- let redfun' = Tacmach.New.pf_apply redfun gl in
+ let redfun' c = Tacmach.New.pf_apply redfun gl c in
match decl with
| LocalAssum (id,ty) ->
if where == InHypValueOnly then
@@ -692,12 +697,12 @@ let bind_red_expr_occurrences occs nbcl redexp =
certain hypothesis *)
let reduct_in_concl (redfun,sty) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
convert_concl_no_check (Tacmach.New.pf_apply redfun gl (Tacmach.New.pf_concl gl)) sty
end }
let reduct_in_hyp ?(check=false) redfun (id,where) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
convert_hyp ~check (pf_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl)
end }
@@ -726,14 +731,14 @@ let pf_e_reduce_decl redfun where decl gl =
Sigma (LocalDef (id, b', ty'), sigma, p +> q)
let e_reduct_in_concl ~check (redfun, sty) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let Sigma (c', sigma, p) = redfun.e_redfun (Tacmach.New.pf_env gl) sigma (Tacmach.New.pf_concl gl) in
Sigma (convert_concl ~check c' sty, sigma, p)
end }
let e_reduct_in_hyp ?(check=false) redfun (id, where) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let Sigma (decl', sigma, p) = pf_e_reduce_decl redfun where (Tacmach.New.pf_get_hyp id gl) gl in
Sigma (convert_hyp ~check decl', sigma, p)
end }
@@ -748,7 +753,7 @@ let e_reduct_option ?(check=false) redfun = function
let e_change_in_concl (redfun,sty) =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
- let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.raw_concl gl) in
+ let Sigma (c, sigma, p) = redfun.e_redfun (Proofview.Goal.env gl) sigma (Proofview.Goal.concl gl) in
Sigma (convert_concl_no_check c sty, sigma, p)
end }
@@ -777,7 +782,7 @@ let e_change_in_hyp redfun (id,where) =
Sigma (convert_hyp c, sigma, p)
end }
-type change_arg = Pattern.patvar_map -> constr Sigma.run
+type change_arg = Pattern.patvar_map -> EConstr.constr Sigma.run
let make_change_arg c pats =
{ run = fun sigma -> Sigma.here (replace_vars (Id.Map.bindings pats) c) sigma }
@@ -791,15 +796,15 @@ let check_types env sigma mayneedglobalcheck deep newc origc =
let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in
if not b then
if
- isSort (whd_all env sigma t1) &&
- isSort (whd_all env sigma t2)
+ isSort sigma (whd_all env sigma t1) &&
+ isSort sigma (whd_all env sigma t2)
then (mayneedglobalcheck := true; sigma)
else
user_err ~hdr:"convert-check-hyp" (str "Types are incompatible.")
else sigma
end
else
- if not (isSort (whd_all env sigma t1)) then
+ if not (isSort sigma (whd_all env sigma t1)) then
user_err ~hdr:"convert-check-hyp" (str "Not a type.")
else sigma
@@ -888,7 +893,7 @@ let reduction_clause redexp cl =
let reduce redexp cl =
let trace () =
let open Printer in
- let pr = (pr_constr, pr_lconstr, pr_evaluable_reference, pr_constr_pattern) in
+ let pr = (pr_econstr, pr_leconstr, pr_evaluable_reference, pr_constr_pattern) in
Pp.(hov 2 (Pputils.pr_red_expr pr str redexp))
in
Proofview.Trace.name_tactic trace begin
@@ -936,13 +941,13 @@ let build_intro_tac id dest tac = match dest with
let rec intro_then_gen name_flag move_flag force_flag dep_flag tac =
let open Context.Rel.Declaration in
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
- let concl = nf_evar (Tacmach.New.project gl) concl in
- match kind_of_term concl with
- | Prod (name,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
+ match EConstr.kind sigma concl with
+ | Prod (name,t,u) when not dep_flag || not (noccurn sigma 1 u) ->
let name = find_name false (LocalAssum (name,t)) name_flag gl in
build_intro_tac name move_flag tac
- | LetIn (name,b,t,u) when not dep_flag || (dependent (mkRel 1) u) ->
+ | LetIn (name,b,t,u) when not dep_flag || not (noccurn sigma 1 u) ->
let name = find_name false (LocalDef (name,b,t)) name_flag gl in
build_intro_tac name move_flag tac
| _ ->
@@ -1067,14 +1072,14 @@ let intros_replacing ids =
(* User-level introduction tactics *)
-let lookup_hypothesis_as_renamed env ccl = function
- | AnonHyp n -> Detyping.lookup_index_as_renamed env ccl n
- | NamedHyp id -> Detyping.lookup_name_as_displayed env ccl id
+let lookup_hypothesis_as_renamed env sigma ccl = function
+ | AnonHyp n -> Detyping.lookup_index_as_renamed env sigma ccl n
+ | NamedHyp id -> Detyping.lookup_name_as_displayed env sigma ccl id
let lookup_hypothesis_as_renamed_gen red h gl =
let env = Proofview.Goal.env gl in
let rec aux ccl =
- match lookup_hypothesis_as_renamed env ccl h with
+ match lookup_hypothesis_as_renamed env (Tacmach.New.project gl) ccl h with
| None when red ->
let (redfun, _) = Redexpr.reduction_of_red_expr env (Red true) in
let Sigma (c, _, _) = redfun.e_redfun env (Proofview.Goal.sigma gl) ccl in
@@ -1107,7 +1112,7 @@ let depth_of_quantified_hypothesis red h gl =
str".")
let intros_until_gen red h =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let n = depth_of_quantified_hypothesis red h gl in
Tacticals.New.tclDO n (if red then introf else intro)
end }
@@ -1217,13 +1222,13 @@ let cut c =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Proofview.Goal.concl gl in
let is_sort =
try
(** Backward compat: ensure that [c] is well-typed. *)
let typ = Typing.unsafe_type_of env sigma c in
let typ = whd_all env sigma typ in
- match kind_of_term typ with
+ match EConstr.kind sigma typ with
| Sort _ -> true
| _ -> false
with e when Pretype_errors.precatchable_exception e -> false
@@ -1243,6 +1248,7 @@ let cut c =
end }
let error_uninstantiated_metas t clenv =
+ let t = EConstr.Unsafe.to_constr t in
let na = meta_name clenv.evd (List.hd (Metaset.elements (metavars_of t))) in
let id = match na with Name id -> id | _ -> anomaly (Pp.str "unnamed dependent meta")
in user_err (str "Cannot find an instance for " ++ pr_id id ++ str".")
@@ -1281,7 +1287,7 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true)
in
let new_hyp_typ = clenv_type clenv in
if not with_evars then check_unresolved_evars_of_metas sigma0 clenv;
- if not with_evars && occur_meta new_hyp_typ then
+ if not with_evars && occur_meta clenv.evd new_hyp_typ then
error_uninstantiated_metas new_hyp_typ clenv;
let new_hyp_prf = clenv_value clenv in
let exact_tac = Proofview.V82.tactic (Tacmach.refine_no_check new_hyp_prf) in
@@ -1300,22 +1306,22 @@ let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true)
(* Elimination tactics *)
(********************************************)
-let last_arg c = match kind_of_term c with
+let last_arg sigma c = match EConstr.kind sigma c with
| App (f,cl) ->
Array.last cl
| _ -> anomaly (Pp.str "last_arg")
-let nth_arg i c =
- if Int.equal i (-1) then last_arg c else
- match kind_of_term c with
+let nth_arg sigma i c =
+ if Int.equal i (-1) then last_arg sigma c else
+ match EConstr.kind sigma c with
| App (f,cl) -> cl.(i)
| _ -> anomaly (Pp.str "nth_arg")
-let index_of_ind_arg t =
- let rec aux i j t = match kind_of_term t with
+let index_of_ind_arg sigma t =
+ let rec aux i j t = match EConstr.kind sigma t with
| Prod (_,t,u) ->
(* heuristic *)
- if isInd (fst (decompose_app t)) then aux (Some j) (j+1) u
+ if isInd sigma (fst (decompose_app sigma t)) then aux (Some j) (j+1) u
else aux i (j+1) u
| _ -> match i with
| Some i -> i
@@ -1330,14 +1336,14 @@ let enforce_prop_bound_names rename tac =
(* so as to avoid having hypothesis such as "t:True", "n:~A" when calling *)
(* elim or induction with schemes built by Indrec.build_induction_scheme *)
let rec aux env sigma i t =
- if i = 0 then t else match kind_of_term t with
+ if i = 0 then t else match EConstr.kind sigma t with
| Prod (Name _ as na,t,t') ->
let very_standard = true in
let na =
if Retyping.get_sort_family_of env sigma t = InProp then
(* "very_standard" says that we should have "H" names only, but
this would break compatibility even more... *)
- let s = match Namegen.head_name t with
+ let s = match Namegen.head_name sigma t with
| Some id when not very_standard -> string_of_id id
| _ -> "" in
Name (add_suffix Namegen.default_prop_ident s)
@@ -1348,9 +1354,9 @@ let enforce_prop_bound_names rename tac =
mkProd (Anonymous,t,aux (push_rel (LocalAssum (Anonymous,t)) env) sigma (i-1) t')
| LetIn (na,c,t,t') ->
mkLetIn (na,c,t,aux (push_rel (LocalDef (na,c,t)) env) sigma (i-1) t')
- | _ -> print_int i; Feedback.msg_notice (print_constr t); assert false in
+ | _ -> assert false in
let rename_branch i =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let t = Proofview.Goal.concl gl in
@@ -1362,10 +1368,10 @@ let enforce_prop_bound_names rename tac =
| _ ->
tac
-let rec contract_letin_in_lam_header c =
- match kind_of_term c with
- | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header c)
- | LetIn (x,b,t,c) -> contract_letin_in_lam_header (subst1 b c)
+let rec contract_letin_in_lam_header sigma c =
+ match EConstr.kind sigma c with
+ | Lambda (x,t,c) -> mkLambda (x,t,contract_letin_in_lam_header sigma c)
+ | LetIn (x,b,t,c) -> contract_letin_in_lam_header sigma (subst1 b c)
| _ -> c
let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags ())
@@ -1373,10 +1379,10 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let elim = contract_letin_in_lam_header elim in
+ let elim = contract_letin_in_lam_header sigma elim in
let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in
let indmv =
- (match kind_of_term (nth_arg i elimclause.templval.rebus) with
+ (match EConstr.kind sigma (nth_arg sigma i elimclause.templval.rebus) with
| Meta mv -> mv
| _ -> user_err ~hdr:"elimination_clause"
(str "The type of elimination clause is not well-formed."))
@@ -1396,7 +1402,7 @@ let elimination_clause_scheme with_evars ?(with_classes=true) ?(flags=elim_flags
type eliminator = {
elimindex : int option; (* None = find it automatically *)
elimrename : (bool * int array) option; (** None = don't rename Prop hyps with H-names *)
- elimbody : constr with_bindings
+ elimbody : EConstr.constr with_bindings
}
let general_elim_clause_gen elimtac indclause elim =
@@ -1406,7 +1412,7 @@ let general_elim_clause_gen elimtac indclause elim =
let (elimc,lbindelimc) = elim.elimbody in
let elimt = Retyping.get_type_of env sigma elimc in
let i =
- match elim.elimindex with None -> index_of_ind_arg elimt | Some i -> i in
+ match elim.elimindex with None -> index_of_ind_arg sigma elimt | Some i -> i in
elimtac elim.elimrename i (elimc, elimt, lbindelimc) indclause
end }
@@ -1428,18 +1434,20 @@ let general_elim with_evars clear_flag (c, lbindc) elim =
(* Case analysis tactics *)
let general_case_analysis_in_context with_evars clear_flag (c,lbindc) =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let concl = Proofview.Goal.concl gl in
let t = Retyping.get_type_of env (Sigma.to_evar_map sigma) c in
let (mind,_) = reduce_to_quantified_ind env (Sigma.to_evar_map sigma) t in
let sort = Tacticals.New.elimination_sort_of_goal gl in
+ let mind = on_snd (fun u -> EInstance.kind (Sigma.to_evar_map sigma) u) mind in
let Sigma (elim, sigma, p) =
- if occur_term c concl then
+ if occur_term (Sigma.to_evar_map sigma) c concl then
build_case_analysis_scheme env sigma mind true sort
else
build_case_analysis_scheme_default env sigma mind sort in
+ let elim = EConstr.of_constr elim in
let tac =
(general_elim with_evars clear_flag (c,lbindc)
{elimindex = None; elimbody = (elim,NoBindings);
@@ -1449,7 +1457,8 @@ let general_case_analysis_in_context with_evars clear_flag (c,lbindc) =
end }
let general_case_analysis with_evars clear_flag (c,lbindc as cx) =
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when lbindc == NoBindings ->
Tacticals.New.tclTHEN (try_intros_until_id_check id)
(general_case_analysis_in_context with_evars clear_flag cx)
@@ -1469,6 +1478,7 @@ let is_nonrec mind = (Global.lookup_mind (fst mind)).mind_finite == Decl_kinds.B
let find_ind_eliminator ind s gl =
let gr = lookup_eliminator ind s in
let evd, c = Tacmach.New.pf_apply Evd.fresh_global gl gr in
+ let c = EConstr.of_constr c in
evd, c
let find_eliminator c gl =
@@ -1502,7 +1512,8 @@ let elim_in_context with_evars clear_flag c = function
| None -> default_elim with_evars clear_flag c
let elim with_evars clear_flag (c,lbindc as cx) elim =
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when lbindc == NoBindings ->
Tacticals.New.tclTHEN (try_intros_until_id_check id)
(elim_in_context with_evars clear_flag cx elim)
@@ -1536,9 +1547,9 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ())
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
- let elim = contract_letin_in_lam_header elim in
+ let elim = contract_letin_in_lam_header sigma elim in
let elimclause = make_clenv_binding env sigma (elim, elimty) bindings in
- let indmv = destMeta (nth_arg i elimclause.templval.rebus) in
+ let indmv = destMeta sigma (nth_arg sigma i elimclause.templval.rebus) in
let hypmv =
try match List.remove Int.equal indmv (clenv_independent elimclause) with
| [a] -> a
@@ -1551,7 +1562,7 @@ let elimination_in_clause_scheme with_evars ?(flags=elim_flags ())
let hypclause = mk_clenv_from_env env sigma (Some 0) (hyp, hyp_typ) in
let elimclause'' = clenv_fchain_in id ~flags hypmv elimclause' hypclause in
let new_hyp_typ = clenv_type elimclause'' in
- if Term.eq_constr hyp_typ new_hyp_typ then
+ if EConstr.eq_constr sigma hyp_typ new_hyp_typ then
user_err ~hdr:"general_rewrite_in"
(str "Nothing to rewrite in " ++ pr_id id ++ str".");
clenv_refine_in with_evars id id sigma elimclause''
@@ -1576,21 +1587,23 @@ let make_projection env sigma params cstr sign elim i n c u =
let elim = match elim with
| NotADefinedRecordUseScheme elim ->
(* bugs: goes from right to left when i increases! *)
- let decl = List.nth cstr.cs_args i in
+ let cs_args = List.map (fun d -> map_rel_decl EConstr.of_constr d) cstr.cs_args in
+ let decl = List.nth cs_args i in
let t = RelDecl.get_type decl in
let b = match decl with LocalAssum _ -> mkRel (i+1) | LocalDef (_,b,_) -> b in
- let branch = it_mkLambda_or_LetIn b cstr.cs_args in
+ let branch = it_mkLambda_or_LetIn b cs_args in
if
(* excludes dependent projection types *)
- noccur_between 1 (n-i-1) t
+ noccur_between sigma 1 (n-i-1) t
(* to avoid surprising unifications, excludes flexible
projection types or lambda which will be instantiated by Meta/Evar *)
- && not (isEvar (fst (whd_betaiota_stack sigma t)))
- && (accept_universal_lemma_under_conjunctions () || not (isRel t))
+ && not (isEvar sigma (fst (whd_betaiota_stack sigma t)))
+ && (accept_universal_lemma_under_conjunctions () || not (isRel sigma t))
then
let t = lift (i+1-n) t in
- let abselim = beta_applist (elim,params@[t;branch]) in
- let c = beta_applist (abselim, [mkApp (c, Context.Rel.to_extended_vect 0 sign)]) in
+ let abselim = beta_applist sigma (elim, params@[t;branch]) in
+ let args = Context.Rel.to_extended_vect mkRel 0 sign in
+ let c = beta_applist sigma (abselim, [mkApp (c, args)]) in
Some (it_mkLambda_or_LetIn c sign, it_mkProd_or_LetIn t sign)
else
None
@@ -1598,7 +1611,7 @@ let make_projection env sigma params cstr sign elim i n c u =
(* goes from left to right when i increases! *)
match List.nth l i with
| Some proj ->
- let args = Context.Rel.to_extended_vect 0 sign in
+ let args = Context.Rel.to_extended_vect mkRel 0 sign in
let proj =
if Environ.is_projection proj env then
mkProj (Projection.make proj false, mkApp (c, args))
@@ -1613,25 +1626,28 @@ let make_projection env sigma params cstr sign elim i n c u =
in elim
let descend_in_conjunctions avoid tac (err, info) c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
try
let t = Retyping.get_type_of env sigma c in
let ((ind,u),t) = reduce_to_quantified_ind env sigma t in
- let sign,ccl = decompose_prod_assum t in
- match match_with_tuple ccl with
+ let sign,ccl = EConstr.decompose_prod_assum sigma t in
+ match match_with_tuple sigma ccl with
| Some (_,_,isrec) ->
let n = (constructors_nrealargs ind).(0) in
let sort = Tacticals.New.elimination_sort_of_goal gl in
let IndType (indf,_) = find_rectype env sigma ccl in
let (_,inst), params = dest_ind_family indf in
+ let params = List.map EConstr.of_constr params in
let cstr = (get_constructors env indf).(0) in
let elim =
try DefinedRecord (Recordops.lookup_projections ind)
with Not_found ->
+ let u = EInstance.kind sigma u in
let sigma = Sigma.Unsafe.of_evar_map sigma in
let Sigma (elim, _, _) = build_case_analysis_scheme env sigma (ind,u) false sort in
+ let elim = EConstr.of_constr elim in
NotADefinedRecordUseScheme elim in
Tacticals.New.tclORELSE0
(Tacticals.New.tclFIRST
@@ -1658,7 +1674,7 @@ let descend_in_conjunctions avoid tac (err, info) c =
(****************************************************)
let solve_remaining_apply_goals =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
if !apply_solve_class_goals then
try
@@ -1682,15 +1698,16 @@ let tclORELSEOPT t k =
Proofview.tclZERO ~info e
| Some tac -> tac)
-let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind)) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind : EConstr.constr with_bindings)) =
+ Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
let flags =
if with_delta then default_unify_flags () else default_no_delta_unify_flags () in
(* The actual type of the theorem. It will be matched against the
goal. If this fails, then the head constant will be unfolded step by
step. *)
- let concl_nprod = nb_prod_modulo_zeta concl in
+ let concl_nprod = nb_prod_modulo_zeta sigma concl in
let rec try_main_apply with_destruct c =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
@@ -1699,7 +1716,7 @@ let general_apply with_delta with_destruct with_evars clear_flag (loc,(c,lbind))
let thm_ty0 = nf_betaiota sigma (Retyping.get_type_of env sigma c) in
let try_apply thm_ty nprod =
try
- let n = nb_prod_modulo_zeta thm_ty - nprod in
+ let n = nb_prod_modulo_zeta sigma thm_ty - nprod in
if n<0 then error "Applied theorem has not enough premisses.";
let clause = make_clenv_binding_apply env sigma (Some n) (c,thm_ty) lbind in
Clenvtac.res_pf clause ~with_evars ~flags
@@ -1826,9 +1843,9 @@ let progress_with_clause flags innerclause clause =
let explain_unable_to_apply_lemma loc env sigma thm innerclause =
user_err ~loc (hov 0
(Pp.str "Unable to apply lemma of type" ++ brk(1,1) ++
- Pp.quote (Printer.pr_lconstr_env env sigma thm) ++ spc() ++
+ Pp.quote (Printer.pr_leconstr_env env sigma thm) ++ spc() ++
str "on hypothesis of type" ++ brk(1,1) ++
- Pp.quote (Printer.pr_lconstr_env innerclause.env innerclause.evd (clenv_type innerclause)) ++
+ Pp.quote (Printer.pr_leconstr_env innerclause.env innerclause.evd (clenv_type innerclause)) ++
str "."))
let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
@@ -1848,7 +1865,7 @@ let apply_in_once_main flags innerclause env sigma (loc,d,lbind) =
let apply_in_once sidecond_first with_delta with_destruct with_evars naming
id (clear_flag,(loc,(d,lbind))) tac =
let open Context.Rel.Declaration in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
let flags =
@@ -1909,9 +1926,10 @@ let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars nam
*)
let cut_and_apply c =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- match kind_of_term (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with
- | Prod (_,c1,c2) when not (dependent (mkRel 1) c2) ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
+ match EConstr.kind sigma (Tacmach.New.pf_hnf_constr gl (Tacmach.New.pf_unsafe_type_of gl c)) with
+ | Prod (_,c1,c2) when Vars.noccurn sigma 1 c2 ->
let concl = Proofview.Goal.concl gl in
let env = Tacmach.New.pf_env gl in
Refine.refine { run = begin fun sigma ->
@@ -1954,7 +1972,7 @@ let exact_check c =
let cast_no_check cast c =
Proofview.Goal.enter { enter = begin fun gl ->
let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
- exact_no_check (Term.mkCast (c, cast, concl))
+ exact_no_check (mkCast (c, cast, concl))
end }
let vm_cast_no_check c = cast_no_check Term.VMcast c
@@ -1962,10 +1980,11 @@ let native_cast_no_check c = cast_no_check Term.NATIVEcast c
let exact_proof c =
let open Tacmach.New in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Refine.refine { run = begin fun sigma ->
let sigma = Sigma.to_evar_map sigma in
let (c, ctx) = Constrintern.interp_casted_constr (pf_env gl) sigma c (pf_concl gl) in
+ let c = EConstr.of_constr c in
let sigma = Evd.merge_universe_context sigma ctx in
Sigma.Unsafe.of_pair (c, sigma)
end }
@@ -1983,7 +2002,7 @@ let assumption =
let concl = Proofview.Goal.concl gl in
let sigma = Tacmach.New.project gl in
let (sigma, is_same_type) =
- if only_eq then (sigma, Constr.equal t concl)
+ if only_eq then (sigma, EConstr.eq_constr sigma t concl)
else
let env = Proofview.Goal.env gl in
infer_conv env sigma t concl
@@ -1997,7 +2016,7 @@ let assumption =
let hyps = Proofview.Goal.hyps gl in
arec gl true hyps
end } in
- Proofview.Goal.nf_enter assumption_tac
+ Proofview.Goal.enter assumption_tac
(*****************************************************************)
(* Modification of a local context *)
@@ -2058,7 +2077,7 @@ let clear_body ids =
(** Do no recheck hypotheses that do not depend *)
let sigma =
if not seen then sigma
- else if List.exists (fun id -> occur_var_in_decl env id decl) ids then
+ else if List.exists (fun id -> occur_var_in_decl env sigma id decl) ids then
check_decl env sigma decl
else sigma
in
@@ -2067,7 +2086,7 @@ let clear_body ids =
in
let (env, sigma, _) = List.fold_left check (base_env, sigma, false) (List.rev ctx) in
let sigma =
- if List.exists (fun id -> occur_var env id concl) ids then
+ if List.exists (fun id -> occur_var env sigma id concl) ids then
check_is_type env sigma concl
else sigma
in
@@ -2102,15 +2121,17 @@ let rec intros_clearing = function
(* Keeping only a few hypotheses *)
let keep hyps =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Proofview.tclENV >>= fun env ->
let ccl = Proofview.Goal.concl gl in
+ let sigma = Tacmach.New.project gl in
let cl,_ =
fold_named_context_reverse (fun (clear,keep) decl ->
+ let decl = map_named_decl EConstr.of_constr decl in
let hyp = NamedDecl.get_id decl in
if Id.List.mem hyp hyps
- || List.exists (occur_var_in_decl env hyp) keep
- || occur_var env hyp ccl
+ || List.exists (occur_var_in_decl env sigma hyp) keep
+ || occur_var env sigma hyp ccl
then (clear,decl::keep)
else (hyp::clear,keep))
~init:([],[]) (Proofview.Goal.env gl)
@@ -2148,9 +2169,9 @@ let bring_hyps hyps =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
let newcl = List.fold_right mkNamedProd_or_LetIn hyps concl in
- let args = Array.of_list (Context.Named.to_instance hyps) in
+ let args = Array.of_list (Context.Named.to_instance mkVar hyps) in
Refine.refine { run = begin fun sigma ->
let Sigma (ev, sigma, p) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
@@ -2182,7 +2203,7 @@ let check_number_of_constructors expctdnumopt i nconstr =
let constructor_tac with_evars expctdnumopt i lbind =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
- let cl = Tacmach.New.pf_nf_concl gl in
+ let cl = Tacmach.New.pf_concl gl in
let reduce_to_quantified_ind =
Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl
in
@@ -2191,9 +2212,9 @@ let constructor_tac with_evars expctdnumopt i lbind =
Array.length (snd (Global.lookup_inductive (fst mind))).mind_consnames in
check_number_of_constructors expctdnumopt i nconstr;
- let Sigma (cons, sigma, p) = Sigma.fresh_constructor_instance
+ let Sigma ((cons, u), sigma, p) = Sigma.fresh_constructor_instance
(Proofview.Goal.env gl) sigma (fst mind, i) in
- let cons = mkConstructU cons in
+ let cons = mkConstructU (cons, EInstance.make u) in
let apply_tac = general_apply true false with_evars None (dloc,(cons,lbind)) in
let tac =
@@ -2221,7 +2242,7 @@ let any_constructor with_evars tacopt =
let t = match tacopt with None -> Proofview.tclUNIT () | Some t -> t in
let tac i = Tacticals.New.tclTHEN (constructor_tac with_evars None i NoBindings) t in
Proofview.Goal.enter { enter = begin fun gl ->
- let cl = Tacmach.New.pf_nf_concl gl in
+ let cl = Tacmach.New.pf_concl gl in
let reduce_to_quantified_ind =
Tacmach.New.pf_apply Tacred.reduce_to_quantified_ind gl
in
@@ -2281,7 +2302,7 @@ let my_find_eq_data_decompose gl t =
| Constr_matching.PatternMatchingFailure -> None
let intro_decomp_eq loc l thin tac id =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let c = mkVar id in
let t = Tacmach.New.pf_unsafe_type_of gl c in
let _,t = Tacmach.New.pf_reduce_to_quantified_ind gl t in
@@ -2319,24 +2340,25 @@ let rewrite_hyp_then assert_style with_evars thin l2r id tac =
List.filter (fun (_,id) -> not (Id.equal id id')) thin in
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let type_of = Tacmach.New.pf_unsafe_type_of gl in
let whd_all = Tacmach.New.pf_apply whd_all gl in
let t = whd_all (type_of (mkVar id)) in
- let eqtac, thin = match match_with_equality_type t with
+ let eqtac, thin = match match_with_equality_type sigma t with
| Some (hdcncl,[_;lhs;rhs]) ->
- if l2r && isVar lhs && not (occur_var env (destVar lhs) rhs) then
- let id' = destVar lhs in
+ if l2r && isVar sigma lhs && not (occur_var env sigma (destVar sigma lhs) rhs) then
+ let id' = destVar sigma lhs in
subst_on l2r id' rhs, early_clear id' thin
- else if not l2r && isVar rhs && not (occur_var env (destVar rhs) lhs) then
- let id' = destVar rhs in
+ else if not l2r && isVar sigma rhs && not (occur_var env sigma (destVar sigma rhs) lhs) then
+ let id' = destVar sigma rhs in
subst_on l2r id' lhs, early_clear id' thin
else
Tacticals.New.tclTHEN (rew_on l2r onConcl) (clear [id]),
thin
| Some (hdcncl,[c]) ->
let l2r = not l2r in (* equality of the form eq_true *)
- if isVar c then
- let id' = destVar c in
+ if isVar sigma c then
+ let id' = destVar sigma c in
Tacticals.New.tclTHEN (rew_on l2r allHypsAndConcl)
(clear_var_and_eq id'),
early_clear id' thin
@@ -2535,9 +2557,9 @@ let ipat_of_name = function
| Anonymous -> None
| Name id -> Some (dloc, IntroNaming (IntroIdentifier id))
-let head_ident c =
- let c = fst (decompose_app ((strip_lam_assum c))) in
- if isVar c then Some (destVar c) else None
+let head_ident sigma c =
+ let c = fst (decompose_app sigma (snd (decompose_lam_assum sigma c))) in
+ if isVar sigma c then Some (destVar sigma c) else None
let assert_as first hd ipat t =
let naming,tac = prepare_intros false IntroAnonymous MoveLast ipat in
@@ -2619,7 +2641,9 @@ let letin_tac_gen with_eq (id,depdecls,lastlhyp,ccl,c) ty =
let eqdata = build_coq_eq_data () in
let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in
let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in
+ let eq = EConstr.of_constr eq in
let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in
+ let refl = EConstr.of_constr refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
let term = mkNamedLetIn id c t (mkLetIn (Name heq, refl, eq, ccl)) in
@@ -2650,6 +2674,7 @@ let insert_before decls lasthyp env =
| Some id ->
Environ.fold_named_context
(fun _ d env ->
+ let d = map_named_decl EConstr.of_constr d in
let env = if Id.equal id (NamedDecl.get_id d) then push_named_context decls env else env in
push_named d env)
~init:(reset_context env) env
@@ -2674,7 +2699,9 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
let eqdata = build_coq_eq_data () in
let args = if lr then [t;mkVar id;c] else [t;c;mkVar id]in
let Sigma (eq, sigma, p) = Sigma.fresh_global env sigma eqdata.eq in
+ let eq = EConstr.of_constr eq in
let Sigma (refl, sigma, q) = Sigma.fresh_global env sigma eqdata.refl in
+ let refl = EConstr.of_constr refl in
let eq = applist (eq,args) in
let refl = applist (refl, [t;mkVar id]) in
let newenv = insert_before [LocalAssum (heq,eq); decl] lastlhyp env in
@@ -2686,7 +2713,7 @@ let mkletin_goal env sigma store with_eq dep (id,lastlhyp,ccl,c) ty =
Sigma (mkNamedLetIn id c t x, sigma, p)
let letin_tac with_eq id c ty occs =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let ccl = Proofview.Goal.concl gl in
@@ -2703,7 +2730,7 @@ let letin_tac with_eq id c ty occs =
end }
let letin_pat_tac with_eq id c occs =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let ccl = Proofview.Goal.concl gl in
@@ -2725,7 +2752,8 @@ let forward b usetac ipat c =
| None ->
Proofview.Goal.enter { enter = begin fun gl ->
let t = Tacmach.New.pf_get_type_of gl c in
- let hd = head_ident c in
+ let sigma = Tacmach.New.project gl in
+ let hd = head_ident sigma c in
Tacticals.New.tclTHENFIRST (assert_as true hd ipat t) (exact_no_check c)
end }
| Some tac ->
@@ -2748,22 +2776,22 @@ let enough_by na t tac = forward false (Some (Some tac)) (ipat_of_name na) t
(* Compute a name for a generalization *)
-let generalized_name c t ids cl = function
+let generalized_name env sigma c t ids cl = function
| Name id as na ->
if Id.List.mem id ids then
user_err (pr_id id ++ str " is already used.");
na
| Anonymous ->
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id ->
(* Keep the name even if not occurring: may be used by intros later *)
Name id
| _ ->
- if noccurn 1 cl then Anonymous else
+ if noccurn sigma 1 cl then Anonymous else
(* On ne s'etait pas casse la tete : on avait pris pour nom de
variable la premiere lettre du type, meme si "c" avait ete une
constante dont on aurait pu prendre directement le nom *)
- named_hd (Global.env()) t Anonymous
+ named_hd env sigma t Anonymous
(* Abstract over [c] in [forall x1:A1(c)..xi:Ai(c).T(c)] producing
[forall x, x1:A1(x1), .., xi:Ai(x). T(x)] with all [c] abtracted in [Ai]
@@ -2771,11 +2799,11 @@ let generalized_name c t ids cl = function
let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl =
let open Context.Rel.Declaration in
- let decls,cl = decompose_prod_n_assum i cl in
+ let decls,cl = decompose_prod_n_assum sigma i cl in
let dummy_prod = it_mkProd_or_LetIn mkProp decls in
- let newdecls,_ = decompose_prod_n_assum i (subst_term_gen eq_constr_nounivs c dummy_prod) in
+ let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in
let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in
- let na = generalized_name c t ids cl' na in
+ let na = generalized_name env sigma c t ids cl' na in
let decl = match b with
| None -> LocalAssum (na,t)
| Some b -> LocalDef (na,b,t)
@@ -2788,13 +2816,20 @@ let generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
let sigma, t = Typing.type_of env sigma c in
generalize_goal_gen env sigma ids i o t cl
+let new_generalize_goal gl i ((occs,c,b),na as o) (cl,sigma) =
+ let env = Tacmach.New.pf_env gl in
+ let ids = Tacmach.New.pf_ids_of_hyps gl in
+ let sigma, t = Typing.type_of env sigma c in
+ generalize_goal_gen env sigma ids i o t cl
+
let old_generalize_dep ?(with_let=false) c gl =
let env = pf_env gl in
let sign = pf_hyps gl in
+ let sigma = project gl in
let init_ids = ids_of_named_context (Global.named_context()) in
- let seek (d:Context.Named.Declaration.t) (toquant:Context.Named.t) =
- if List.exists (fun d' -> occur_var_in_decl env (NamedDecl.get_id d') d) toquant
- || dependent_in_decl c d then
+ let seek (d:named_declaration) (toquant:named_context) =
+ if List.exists (fun d' -> occur_var_in_decl env sigma (NamedDecl.get_id d') d) toquant
+ || dependent_in_decl sigma c d then
d::toquant
else
toquant in
@@ -2803,7 +2838,7 @@ let old_generalize_dep ?(with_let=false) c gl =
let qhyps = List.map NamedDecl.get_id to_quantify_rev in
let tothin = List.filter (fun id -> not (Id.List.mem id init_ids)) qhyps in
let tothin' =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id when mem_named_context_val id (val_of_named_context sign) && not (Id.List.mem id init_ids)
-> id::tothin
| _ -> tothin
@@ -2811,7 +2846,7 @@ let old_generalize_dep ?(with_let=false) c gl =
let cl' = it_mkNamedProd_or_LetIn (Tacmach.pf_concl gl) to_quantify in
let body =
if with_let then
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id -> id |> Tacmach.pf_get_hyp gl |> NamedDecl.get_value
| _ -> None
else None
@@ -2820,7 +2855,7 @@ let old_generalize_dep ?(with_let=false) c gl =
(cl',project gl) in
(** Check that the generalization is indeed well-typed *)
let (evd, _) = Typing.type_of env evd cl'' in
- let args = Context.Named.to_instance to_quantify_rev in
+ let args = Context.Named.to_instance mkVar to_quantify_rev in
tclTHENLIST
[tclEVARS evd;
Proofview.V82.of_tactic (apply_type cl'' (if Option.is_empty body then c::args else args));
@@ -2831,10 +2866,10 @@ let generalize_dep ?(with_let = false) c =
Proofview.V82.tactic (old_generalize_dep ~with_let c)
(** *)
-let generalize_gen_let lconstr = Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+let generalize_gen_let lconstr = Proofview.Goal.s_enter { s_enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let newcl, evd =
- List.fold_right_i (Tacmach.New.of_old generalize_goal gl) 0 lconstr
+ List.fold_right_i (new_generalize_goal gl) 0 lconstr
(Tacmach.New.pf_concl gl,Tacmach.New.project gl)
in
let (evd, _) = Typing.type_of env evd newcl in
@@ -2910,19 +2945,19 @@ let specialize (c,lbind) ipat =
let (thd,tstack) = whd_nored_stack clause.evd (clenv_value clause) in
let rec chk = function
| [] -> []
- | t::l -> if occur_meta t then [] else t :: chk l
+ | t::l -> if occur_meta clause.evd t then [] else t :: chk l
in
let tstack = chk tstack in
let term = applist(thd,List.map (nf_evar clause.evd) tstack) in
- if occur_meta term then
+ if occur_meta clause.evd term then
user_err (str "Cannot infer an instance for " ++
- pr_name (meta_name clause.evd (List.hd (collect_metas term))) ++
+ pr_name (meta_name clause.evd (List.hd (collect_metas clause.evd term))) ++
str ".");
clause.evd, term in
let typ = Retyping.get_type_of env sigma term in
let tac =
- match kind_of_term (fst(decompose_app (snd(decompose_lam_assum c)))) with
+ match EConstr.kind sigma (fst(EConstr.decompose_app sigma (snd(EConstr.decompose_lam_assum sigma c)))) with
| Var id when Id.List.mem id (Tacmach.New.pf_ids_of_hyps gl) ->
(* Like assert (id:=id args) but with the concept of specialization *)
let naming,tac =
@@ -2964,6 +2999,7 @@ let unfold_body x =
(pr_id x ++ str" is not a defined hypothesis.")
| LocalDef (_,xval,_) -> xval
in
+ let xval = EConstr.of_constr xval in
Tacticals.New.afterHyp x begin fun aft ->
let hl = List.fold_right (fun decl cl -> (NamedDecl.get_id decl, InHyp) :: cl) aft [] in
let rfun _ _ c = replace_vars [x, xval] c in
@@ -3016,7 +3052,7 @@ let warn_unused_intro_pattern =
strbrk"Unused introduction " ++ str (String.plural (List.length names) "pattern")
++ str": " ++ prlist_with_sep spc
(Miscprint.pr_intro_pattern
- (fun c -> Printer.pr_constr (fst (run_delayed (Global.env()) Evd.empty c)))) names)
+ (fun c -> Printer.pr_econstr (fst (run_delayed (Global.env()) Evd.empty c)))) names)
let check_unused_names names =
if not (List.is_empty names) then
@@ -3150,12 +3186,12 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names =
substitutions aussi sur l'argument voisin *)
let expand_projections env sigma c =
- let sigma = Sigma.to_evar_map sigma in
let rec aux env c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Proj (p, c) -> Retyping.expand_projection env sigma p (aux env c) []
- | _ -> map_constr_with_full_binders push_rel aux env c
- in aux env c
+ | _ -> map_constr_with_full_binders sigma push_rel aux env c
+ in
+ aux env c
(* Marche pas... faut prendre en compte l'occurrence précise... *)
@@ -3163,13 +3199,13 @@ let expand_projections env sigma c =
let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
+ let sigma = Tacmach.New.project gl in
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 (Proofview.Goal.assume gl) in
let reduce_to_quantified_ref = Tacmach.New.pf_apply reduce_to_quantified_ref gl in
let typ0 = reduce_to_quantified_ref indref tmptyp0 in
- let prods, indtyp = decompose_prod_assum typ0 in
- let hd,argl = decompose_app indtyp in
+ let prods, indtyp = decompose_prod_assum sigma typ0 in
+ let hd,argl = decompose_app sigma indtyp in
let env' = push_rel_context prods env in
- let sigma = Proofview.Goal.sigma gl in
let params = List.firstn nparams argl in
let params' = List.map (expand_projections env' sigma) params in
(* le gl est important pour ne pas préévaluer *)
@@ -3181,17 +3217,18 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
(tac avoid)
else
let c = List.nth argl (i-1) in
- match kind_of_term c with
- | Var id when not (List.exists (occur_var env id) args') &&
- not (List.exists (occur_var env id) params') ->
+ match EConstr.kind sigma c with
+ | Var id when not (List.exists (fun c -> occur_var env sigma id c) args') &&
+ not (List.exists (fun c -> occur_var env sigma id c) params') ->
(* Based on the knowledge given by the user, all
constraints on the variable are generalizable in the
current environment so that it is clearable after destruction *)
atomize_one (i-1) (c::args) (c::args') (id::avoid)
| _ ->
let c' = expand_projections env' sigma c in
- if List.exists (dependent c) params' ||
- List.exists (dependent c) args'
+ let dependent t = dependent sigma c t in
+ if List.exists dependent params' ||
+ List.exists dependent args'
then
(* This is a case where the argument is constrained in a
way which would require some kind of inversion; we
@@ -3203,11 +3240,11 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
(* We reason blindly on the term and do as if it were
generalizable, ignoring the constraints coming from
its structure *)
- let id = match kind_of_term c with
+ let id = match EConstr.kind sigma c with
| Var id -> id
| _ ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
- id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in
+ id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in
let x = fresh_id_in_env avoid id env in
Tacticals.New.tclTHEN
(letin_tac None (Name x) c None allHypsAndConcl)
@@ -3281,7 +3318,7 @@ let atomize_param_of_ind_then (indref,nparams,_) hyp0 tac =
exception Shunt of Id.t move_location
-let cook_sign hyp0_opt inhyps indvars env =
+let cook_sign hyp0_opt inhyps indvars env sigma =
(* First phase from L to R: get [toclear], [decldep] and [statuslist]
for the hypotheses before (= more ancient than) hyp0 (see above) *)
let toclear = ref [] in
@@ -3293,6 +3330,7 @@ let cook_sign hyp0_opt inhyps indvars env =
let before = ref true in
let maindep = ref false in
let seek_deps env decl rhyp =
+ let decl = map_named_decl EConstr.of_constr decl in
let hyp = NamedDecl.get_id decl in
if (match hyp0_opt with Some hyp0 -> Id.equal hyp hyp0 | _ -> false)
then begin
@@ -3308,11 +3346,11 @@ let cook_sign hyp0_opt inhyps indvars env =
rhyp
end else
let dephyp0 = List.is_empty inhyps &&
- (Option.cata (fun id -> occur_var_in_decl env id decl) false hyp0_opt)
+ (Option.cata (fun id -> occur_var_in_decl env sigma id decl) false hyp0_opt)
in
let depother = List.is_empty inhyps &&
- (List.exists (fun id -> occur_var_in_decl env id decl) indvars ||
- List.exists (fun decl' -> occur_var_in_decl env (NamedDecl.get_id decl') decl) !decldeps)
+ (List.exists (fun id -> occur_var_in_decl env sigma id decl) indvars ||
+ List.exists (fun decl' -> occur_var_in_decl env sigma (NamedDecl.get_id decl') decl) !decldeps)
in
if not (List.is_empty inhyps) && Id.List.mem hyp inhyps
|| dephyp0 || depother
@@ -3385,15 +3423,15 @@ type elim_scheme = {
elimc: constr with_bindings option;
elimt: types;
indref: global_reference option;
- params: Context.Rel.t; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
+ params: rel_context; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
nparams: int; (* number of parameters *)
- predicates: Context.Rel.t; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
+ predicates: rel_context; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
npredicates: int; (* Number of predicates *)
- branches: Context.Rel.t; (* branchr,...,branch1 *)
+ branches: rel_context; (* branchr,...,branch1 *)
nbranches: int; (* Number of branches *)
- args: Context.Rel.t; (* (xni, Ti_ni) ... (x1, Ti_1) *)
+ args: rel_context; (* (xni, Ti_ni) ... (x1, Ti_1) *)
nargs: int; (* number of arguments *)
- indarg: Context.Rel.Declaration.t option; (* Some (H,I prm1..prmp x1...xni)
+ indarg: rel_declaration option; (* Some (H,I prm1..prmp x1...xni)
if HI is in premisses, None otherwise *)
concl: types; (* Qi x1...xni HI (f...), HI and (f...)
are optional and mutually exclusive *)
@@ -3458,13 +3496,13 @@ let error_ind_scheme s =
let s = if not (String.is_empty s) then s^" " else s in
user_err ~hdr:"Tactics" (str "Cannot recognize " ++ str s ++ str "an induction scheme.")
-let glob = Universes.constr_of_global
+let glob c = EConstr.of_constr (Universes.constr_of_global c)
let coq_eq = lazy (glob (Coqlib.build_coq_eq ()))
let coq_eq_refl = lazy (glob (Coqlib.build_coq_eq_refl ()))
-let coq_heq = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq")
-let coq_heq_refl = lazy (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl")
+let coq_heq = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq"))
+let coq_heq_refl = lazy (EConstr.of_constr (Coqlib.coq_constant "mkHEq" ["Logic";"JMeq"] "JMeq_refl"))
let mkEq t x y =
@@ -3491,24 +3529,24 @@ let lift_togethern n l =
let lift_list l = List.map (lift 1) l
-let ids_of_constr ?(all=false) vars c =
+let ids_of_constr sigma ?(all=false) vars c =
let rec aux vars c =
- match kind_of_term c with
+ match EConstr.kind sigma c with
| Var id -> Id.Set.add id vars
| App (f, args) ->
- (match kind_of_term f with
+ (match EConstr.kind sigma f with
| Construct ((ind,_),_)
| Ind (ind,_) ->
let (mib,mip) = Global.lookup_inductive ind in
Array.fold_left_from
(if all then 0 else mib.Declarations.mind_nparams)
aux vars args
- | _ -> Term.fold_constr aux vars c)
- | _ -> Term.fold_constr aux vars c
+ | _ -> EConstr.fold sigma aux vars c)
+ | _ -> EConstr.fold sigma aux vars c
in aux vars c
-let decompose_indapp f args =
- match kind_of_term f with
+let decompose_indapp sigma f args =
+ match EConstr.kind sigma f with
| Construct ((ind,_),_)
| Ind (ind,_) ->
let (mib,mip) = Global.lookup_inductive ind in
@@ -3558,7 +3596,7 @@ let make_abstract_generalize env id typ concl dep ctx body c eqs args refls =
Sigma (mkApp (appeqs, abshypt), sigma, p)
end }
-let hyps_of_vars env sign nogen hyps =
+let hyps_of_vars env sigma sign nogen hyps =
if Id.Set.is_empty hyps then []
else
let (_,lh) =
@@ -3568,7 +3606,7 @@ let hyps_of_vars env sign nogen hyps =
if Id.Set.mem x nogen then (hs,hl)
else if Id.Set.mem x hs then (hs,x::hl)
else
- let xvars = global_vars_set_of_decl env d in
+ let xvars = global_vars_set_of_decl env sigma d in
if not (Id.Set.is_empty (Id.Set.diff xvars hs)) then
(Id.Set.add x hs, x :: hl)
else (hs, hl))
@@ -3578,11 +3616,11 @@ let hyps_of_vars env sign nogen hyps =
exception Seen
-let linear vars args =
+let linear sigma vars args =
let seen = ref vars in
try
Array.iter (fun i ->
- let rels = ids_of_constr ~all:true Id.Set.empty i in
+ let rels = ids_of_constr sigma ~all:true Id.Set.empty i in
let seen' =
Id.Set.fold (fun id acc ->
if Id.Set.mem id acc then raise Seen
@@ -3597,14 +3635,15 @@ let is_defined_variable env id =
env |> lookup_named id |> is_local_def
let abstract_args gl generalize_vars dep id defined f args =
+ let open Tacmach.New in
let open Context.Rel.Declaration in
- let sigma = ref (Tacmach.project gl) in
- let env = Tacmach.pf_env gl in
- let concl = Tacmach.pf_concl gl in
- let dep = dep || dependent (mkVar id) concl in
+ let sigma = ref (Tacmach.New.project gl) in
+ let env = Tacmach.New.pf_env gl in
+ let concl = Tacmach.New.pf_concl gl in
+ let dep = dep || local_occur_var !sigma id concl in
let avoid = ref [] in
let get_id name =
- let id = fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in
+ let id = new_fresh_id !avoid (match name with Name n -> n | Anonymous -> Id.of_string "gen_x") gl in
avoid := id :: !avoid; id
in
(* Build application generalized w.r.t. the argument plus the necessary eqs.
@@ -3619,13 +3658,13 @@ let abstract_args gl generalize_vars dep id defined f args =
let decl = List.hd rel in
RelDecl.get_name decl, RelDecl.get_type decl, c
in
- let argty = Tacmach.pf_unsafe_type_of gl arg in
+ let argty = Tacmach.New.pf_unsafe_type_of gl arg in
let sigma', ty = Evarsolve.refresh_universes (Some true) env !sigma ty in
let () = sigma := sigma' in
let lenctx = List.length ctx in
let liftargty = lift lenctx argty in
- let leq = constr_cmp Reduction.CUMUL liftargty ty in
- match kind_of_term arg with
+ let leq = constr_cmp !sigma Reduction.CUMUL liftargty ty in
+ match EConstr.kind !sigma arg with
| Var id when not (is_defined_variable env id) && leq && not (Id.Set.mem id nongenvars) ->
(subst1 arg arity, ctx, ctxenv, mkApp (c, [|arg|]), args, eqs, refls,
Id.Set.add id nongenvars, Id.Set.remove id vars, env)
@@ -3644,23 +3683,23 @@ let abstract_args gl generalize_vars dep id defined f args =
in
let eqs = eq :: lift_list eqs in
let refls = refl :: refls in
- let argvars = ids_of_constr vars arg in
+ let argvars = ids_of_constr !sigma vars arg in
(arity, ctx, push_rel decl ctxenv, c', args, eqs, refls,
nongenvars, Id.Set.union argvars vars, env)
in
- let f', args' = decompose_indapp f args in
+ let f', args' = decompose_indapp !sigma f args in
let dogen, f', args' =
- let parvars = ids_of_constr ~all:true Id.Set.empty f' in
- if not (linear parvars args') then true, f, args
+ let parvars = ids_of_constr !sigma ~all:true Id.Set.empty f' in
+ if not (linear !sigma parvars args') then true, f, args
else
- match Array.findi (fun i x -> not (isVar x) || is_defined_variable env (destVar x)) args' with
+ match Array.findi (fun i x -> not (isVar !sigma x) || is_defined_variable env (destVar !sigma x)) args' with
| None -> false, f', args'
| Some nonvar ->
let before, after = Array.chop nonvar args' in
true, mkApp (f', before), after
in
if dogen then
- let tyf' = Tacmach.pf_unsafe_type_of gl f' in
+ let tyf' = Tacmach.New.pf_unsafe_type_of gl f' in
let arity, ctx, ctxenv, c', args, eqs, refls, nogen, vars, env =
Array.fold_left aux (tyf',[],env,f',[],[],[],Id.Set.empty,Id.Set.empty,env) args'
in
@@ -3668,14 +3707,14 @@ let abstract_args gl generalize_vars dep id defined f args =
let vars =
if generalize_vars then
let nogen = Id.Set.add id nogen in
- hyps_of_vars (pf_env gl) (pf_hyps gl) nogen vars
+ hyps_of_vars (pf_env gl) (project gl) (Proofview.Goal.hyps gl) nogen vars
else []
in
let body, c' =
if defined then Some c', Retyping.get_type_of ctxenv !sigma c'
else None, c'
in
- let typ = Tacmach.pf_get_hyp_typ gl id in
+ let typ = Tacmach.New.pf_get_hyp_typ id gl in
let tac = make_abstract_generalize (pf_env gl) id typ concl dep ctx body c' eqs args refls in
let tac = Proofview.Unsafe.tclEVARS !sigma <*> tac in
Some (tac, dep, succ (List.length ctx), vars)
@@ -3683,21 +3722,22 @@ let abstract_args gl generalize_vars dep id defined f args =
let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
let open Context.Named.Declaration in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
Coqlib.check_required_library Coqlib.jmeq_module_name;
+ let sigma = Tacmach.New.project gl in
let (f, args, def, id, oldid) =
let oldid = Tacmach.New.pf_get_new_id id gl in
match Tacmach.New.pf_get_hyp id gl with
- | LocalAssum (_,t) -> let f, args = decompose_app t in
+ | LocalAssum (_,t) -> let f, args = decompose_app sigma t in
(f, args, false, id, oldid)
| LocalDef (_,t,_) ->
- let f, args = decompose_app t in
+ let f, args = decompose_app sigma t in
(f, args, true, id, oldid)
in
if List.is_empty args then Proofview.tclUNIT ()
else
let args = Array.of_list args in
- let newc = Tacmach.New.of_old (fun gl -> abstract_args gl generalize_vars force_dep id def f args) gl in
+ let newc = abstract_args gl generalize_vars force_dep id def f args in
match newc with
| None -> Proofview.tclUNIT ()
| Some (tac, dep, n, vars) ->
@@ -3720,9 +3760,12 @@ let abstract_generalize ?(generalize_vars=true) ?(force_dep=false) id =
Tacticals.New.tclTRY (generalize_dep ~with_let:true (mkVar id))) vars])
end }
-let rec compare_upto_variables x y =
- if (isVar x || isRel x) && (isVar y || isRel y) then true
- else compare_constr compare_upto_variables x y
+let compare_upto_variables sigma x y =
+ let rec compare x y =
+ if (isVar sigma x || isRel sigma x) && (isVar sigma y || isRel sigma y) then true
+ else compare_constr sigma compare x y
+ in
+ compare x y
let specialize_eqs id gl =
let open Context.Rel.Declaration in
@@ -3730,21 +3773,21 @@ let specialize_eqs id gl =
let ty = Tacmach.pf_get_hyp_typ gl id in
let evars = ref (project gl) in
let unif env evars c1 c2 =
- compare_upto_variables c1 c2 && Evarconv.e_conv env evars c1 c2
+ compare_upto_variables !evars c1 c2 && Evarconv.e_conv env evars c1 c2
in
let rec aux in_eqs ctx acc ty =
- match kind_of_term ty with
+ match EConstr.kind !evars ty with
| Prod (na, t, b) ->
- (match kind_of_term t with
- | App (eq, [| eqty; x; y |]) when Term.eq_constr (Lazy.force coq_eq) eq ->
- let c = if noccur_between 1 (List.length ctx) x then y else x in
+ (match EConstr.kind !evars t with
+ | App (eq, [| eqty; x; y |]) when EConstr.eq_constr !evars (Lazy.force coq_eq) eq ->
+ let c = if noccur_between !evars 1 (List.length ctx) x then y else x in
let pt = mkApp (Lazy.force coq_eq, [| eqty; c; c |]) in
let p = mkApp (Lazy.force coq_eq_refl, [| eqty; c |]) in
if unif (push_rel_context ctx env) evars pt t then
aux true ctx (mkApp (acc, [| p |])) (subst1 p b)
else acc, in_eqs, ctx, ty
- | App (heq, [| eqty; x; eqty'; y |]) when Term.eq_constr heq (Lazy.force coq_heq) ->
- let eqt, c = if noccur_between 1 (List.length ctx) x then eqty', y else eqty, x in
+ | App (heq, [| eqty; x; eqty'; y |]) when EConstr.eq_constr !evars heq (Lazy.force coq_heq) ->
+ let eqt, c = if noccur_between !evars 1 (List.length ctx) x then eqty', y else eqty, x in
let pt = mkApp (Lazy.force coq_heq, [| eqt; c; eqt; c |]) in
let p = mkApp (Lazy.force coq_heq_refl, [| eqt; c |]) in
if unif (push_rel_context ctx env) evars pt t then
@@ -3760,7 +3803,7 @@ let specialize_eqs id gl =
let acc, worked, ctx, ty = aux false [] (mkVar id) ty in
let ctx' = nf_rel_context_evar !evars ctx in
let ctx'' = List.map (function
- | LocalDef (n,k,t) when isEvar k -> LocalAssum (n,t)
+ | LocalDef (n,k,t) when isEvar !evars k -> LocalAssum (n,t)
| decl -> decl) ctx'
in
let ty' = it_mkProd_or_LetIn ty ctx'' in
@@ -3774,15 +3817,15 @@ let specialize_eqs id gl =
else tclFAIL 0 (str "Nothing to do in hypothesis " ++ pr_id id) gl
-let specialize_eqs id = Proofview.Goal.nf_enter { enter = begin fun gl ->
+let specialize_eqs id = Proofview.Goal.enter { enter = begin fun gl ->
let msg = str "Specialization not allowed on dependent hypotheses" in
Proofview.tclOR (clear [id])
(fun _ -> Tacticals.New.tclZEROMSG msg) >>= fun () ->
Proofview.V82.tactic (specialize_eqs id)
end }
-let occur_rel n c =
- let res = not (noccurn n c) in
+let occur_rel sigma n c =
+ let res = not (noccurn sigma n c) in
res
(* This function splits the products of the induction scheme [elimt] into four
@@ -3793,20 +3836,20 @@ let occur_rel n c =
if there is no branch, we try to fill in acc3 with args/indargs.
We also return the conclusion.
*)
-let decompose_paramspred_branch_args elimt =
+let decompose_paramspred_branch_args sigma elimt =
let open Context.Rel.Declaration in
let rec cut_noccur elimt acc2 =
- match kind_of_term elimt with
+ match EConstr.kind sigma elimt with
| Prod(nme,tpe,elimt') ->
- let hd_tpe,_ = decompose_app ((strip_prod_assum tpe)) in
- if not (occur_rel 1 elimt') && isRel hd_tpe
+ let hd_tpe,_ = decompose_app sigma (snd (decompose_prod_assum sigma tpe)) in
+ if not (occur_rel sigma 1 elimt') && isRel sigma hd_tpe
then cut_noccur elimt' (LocalAssum (nme,tpe)::acc2)
- else let acc3,ccl = decompose_prod_assum elimt in acc2 , acc3 , ccl
+ else let acc3,ccl = decompose_prod_assum sigma elimt in acc2 , acc3 , ccl
| App(_, _) | Rel _ -> acc2 , [] , elimt
| _ -> error_ind_scheme "" in
let rec cut_occur elimt acc1 =
- match kind_of_term elimt with
- | Prod(nme,tpe,c) when occur_rel 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1)
+ match EConstr.kind sigma elimt with
+ | Prod(nme,tpe,c) when occur_rel sigma 1 c -> cut_occur c (LocalAssum (nme,tpe)::acc1)
| Prod(nme,tpe,c) -> let acc2,acc3,ccl = cut_noccur elimt [] in acc1,acc2,acc3,ccl
| App(_, _) | Rel _ -> acc1,[],[],elimt
| _ -> error_ind_scheme "" in
@@ -3819,17 +3862,17 @@ let decompose_paramspred_branch_args elimt =
args. We suppose there is only one predicate here. *)
match acc2 with
| [] ->
- let hyps,ccl = decompose_prod_assum elimt in
- let hd_ccl_pred,_ = decompose_app ccl in
- begin match kind_of_term hd_ccl_pred with
+ let hyps,ccl = decompose_prod_assum sigma elimt in
+ let hd_ccl_pred,_ = decompose_app sigma ccl in
+ begin match EConstr.kind sigma hd_ccl_pred with
| Rel i -> let acc3,acc1 = List.chop (i-1) hyps in acc1 , [] , acc3 , ccl
| _ -> error_ind_scheme ""
end
| _ -> acc1, acc2 , acc3, ccl
-let exchange_hd_app subst_hd t =
- let hd,args= decompose_app t in mkApp (subst_hd,Array.of_list args)
+let exchange_hd_app sigma subst_hd t =
+ let hd,args= decompose_app sigma t in mkApp (subst_hd,Array.of_list args)
(* Builds an elim_scheme from its type and calling form (const+binding). We
first separate branches. We obtain branches, hyps before (params + preds),
@@ -3847,14 +3890,14 @@ let exchange_hd_app subst_hd t =
predicates are cited in the conclusion.
- finish to fill in the elim_scheme: indarg/farg/args and finally indref. *)
-let compute_elim_sig ?elimc elimt =
+let compute_elim_sig sigma ?elimc elimt =
let open Context.Rel.Declaration in
let params_preds,branches,args_indargs,conclusion =
- decompose_paramspred_branch_args elimt in
+ decompose_paramspred_branch_args sigma elimt in
- let ccl = exchange_hd_app (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in
+ let ccl = exchange_hd_app sigma (mkVar (Id.of_string "__QI_DUMMY__")) conclusion in
let concl_with_args = it_mkProd_or_LetIn ccl args_indargs in
- let nparams = Int.Set.cardinal (free_rels concl_with_args) in
+ let nparams = Int.Set.cardinal (free_rels sigma concl_with_args) in
let preds,params = List.chop (List.length params_preds - nparams) params_preds in
(* A first approximation, further analysis will tweak it *)
@@ -3863,7 +3906,7 @@ let compute_elim_sig ?elimc elimt =
elimc = elimc; elimt = elimt; concl = conclusion;
predicates = preds; npredicates = List.length preds;
branches = branches; nbranches = List.length branches;
- farg_in_concl = isApp ccl && isApp (last_arg ccl);
+ farg_in_concl = isApp sigma ccl && isApp sigma (last_arg sigma ccl);
params = params; nparams = nparams;
(* all other fields are unsure at this point. Including these:*)
args = args_indargs; nargs = List.length args_indargs; } in
@@ -3884,9 +3927,9 @@ let compute_elim_sig ?elimc elimt =
match List.hd args_indargs with
| LocalDef (hiname,_,hi) -> error_ind_scheme ""
| LocalAssum (hiname,hi) ->
- let hi_ind, hi_args = decompose_app hi in
+ let hi_ind, hi_args = decompose_app sigma hi in
let hi_is_ind = (* hi est d'un type globalisable *)
- match kind_of_term hi_ind with
+ match EConstr.kind sigma hi_ind with
| Ind (mind,_) -> true
| Var _ -> true
| Const _ -> true
@@ -3899,7 +3942,7 @@ let compute_elim_sig ?elimc elimt =
else (* Last arg is the indarg *)
res := {!res with
indarg = Some (List.hd !res.args);
- indarg_in_concl = occur_rel 1 ccl;
+ indarg_in_concl = occur_rel sigma 1 ccl;
args = List.tl !res.args; nargs = !res.nargs - 1;
};
raise Exit);
@@ -3909,49 +3952,49 @@ let compute_elim_sig ?elimc elimt =
| None -> !res (* No indref *)
| Some (LocalDef _) -> error_ind_scheme ""
| Some (LocalAssum (_,ind)) ->
- let indhd,indargs = decompose_app ind in
- try {!res with indref = Some (global_of_constr indhd) }
+ let indhd,indargs = decompose_app sigma ind in
+ try {!res with indref = Some (fst (Termops.global_of_constr sigma indhd)) }
with e when CErrors.noncritical e ->
error "Cannot find the inductive type of the inductive scheme."
-let compute_scheme_signature scheme names_info ind_type_guess =
+let compute_scheme_signature evd scheme names_info ind_type_guess =
let open Context.Rel.Declaration in
- let f,l = decompose_app scheme.concl in
+ let f,l = decompose_app evd scheme.concl in
(* Vérifier que les arguments de Qi sont bien les xi. *)
let cond, check_concl =
match scheme.indarg with
| Some (LocalDef _) ->
error "Strange letin, cannot recognize an induction scheme."
| None -> (* Non standard scheme *)
- let cond hd = Term.eq_constr hd ind_type_guess && not scheme.farg_in_concl
+ let cond hd = EConstr.eq_constr evd hd ind_type_guess && not scheme.farg_in_concl
in (cond, fun _ _ -> ())
| Some (LocalAssum (_,ind)) -> (* Standard scheme from an inductive type *)
- let indhd,indargs = decompose_app ind in
- let cond hd = Term.eq_constr hd indhd in
+ let indhd,indargs = decompose_app evd ind in
+ let cond hd = EConstr.eq_constr evd hd indhd in
let check_concl is_pred p =
(* Check again conclusion *)
let ccl_arg_ok = is_pred (p + scheme.nargs + 1) f == IndArg in
let ind_is_ok =
- List.equal Term.eq_constr
+ List.equal (fun c1 c2 -> EConstr.eq_constr evd c1 c2)
(List.lastn scheme.nargs indargs)
- (Context.Rel.to_extended_list 0 scheme.args) in
+ (Context.Rel.to_extended_list mkRel 0 scheme.args) in
if not (ccl_arg_ok && ind_is_ok) then
error_ind_scheme "the conclusion of"
in (cond, check_concl)
in
let is_pred n c =
- let hd = fst (decompose_app c) in
- match kind_of_term hd with
+ let hd = fst (decompose_app evd c) in
+ match EConstr.kind evd hd with
| Rel q when n < q && q <= n+scheme.npredicates -> IndArg
| _ when cond hd -> RecArg
| _ -> OtherArg
in
let rec check_branch p c =
- match kind_of_term c with
+ match EConstr.kind evd c with
| Prod (_,t,c) ->
- (is_pred p t, true, dependent (mkRel 1) c) :: check_branch (p+1) c
+ (is_pred p t, true, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c
| LetIn (_,_,_,c) ->
- (OtherArg, false, dependent (mkRel 1) c) :: check_branch (p+1) c
+ (OtherArg, false, not (Vars.noccurn evd 1 c)) :: check_branch (p+1) c
| _ when is_pred p c == IndArg -> []
| _ -> raise Exit
in
@@ -3983,55 +4026,61 @@ let compute_scheme_signature scheme names_info ind_type_guess =
the non standard case, naming of generated hypos is slightly
different. *)
let compute_elim_signature (evd,(elimc,elimt),ind_type_guess) names_info =
- let scheme = compute_elim_sig ~elimc:elimc elimt in
- evd, (compute_scheme_signature scheme names_info ind_type_guess, scheme)
+ let scheme = compute_elim_sig evd ~elimc:elimc elimt in
+ evd, (compute_scheme_signature evd scheme names_info ind_type_guess, scheme)
let guess_elim isrec dep s hyp0 gl =
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in
- let mind,_ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in
+ let (mind, u), _ = Tacmach.New.pf_reduce_to_quantified_ind gl tmptyp0 in
let evd, elimc =
- if isrec && not (is_nonrec (fst mind)) then find_ind_eliminator (fst mind) s gl
+ if isrec && not (is_nonrec mind) then find_ind_eliminator mind s gl
else
let env = Tacmach.New.pf_env gl in
let sigma = Sigma.Unsafe.of_evar_map (Tacmach.New.project gl) in
+ let u = EInstance.kind (Tacmach.New.project gl) u in
if use_dependent_propositions_elimination () && dep
then
- let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma mind true s in
+ let Sigma (ind, sigma, _) = build_case_analysis_scheme env sigma (mind, u) true s in
+ let ind = EConstr.of_constr ind in
(Sigma.to_evar_map sigma, ind)
else
- let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma mind s in
+ let Sigma (ind, sigma, _) = build_case_analysis_scheme_default env sigma (mind, u) s in
+ let ind = EConstr.of_constr ind in
(Sigma.to_evar_map sigma, ind)
in
let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in
- evd, ((elimc, NoBindings), elimt), mkIndU mind
+ evd, ((elimc, NoBindings), elimt), mkIndU (mind, u)
let given_elim hyp0 (elimc,lbind as e) gl =
+ let sigma = Tacmach.New.project gl in
let tmptyp0 = Tacmach.New.pf_get_hyp_typ hyp0 gl in
- let ind_type_guess,_ = decompose_app ((strip_prod tmptyp0)) in
- Tacmach.New.project gl, (e, Tacmach.New.pf_unsafe_type_of gl elimc), ind_type_guess
+ let ind_type_guess,_ = decompose_app sigma (snd (decompose_prod sigma tmptyp0)) in
+ let elimt = Tacmach.New.pf_unsafe_type_of gl elimc in
+ Tacmach.New.project gl, (e, elimt), ind_type_guess
type scheme_signature =
(Id.t list * (elim_arg_kind * bool * bool * Id.t) list) array
type eliminator_source =
- | ElimUsing of (eliminator * types) * scheme_signature
+ | ElimUsing of (eliminator * EConstr.types) * scheme_signature
| ElimOver of bool * Id.t
let find_induction_type isrec elim hyp0 gl =
+ let sigma = Tacmach.New.project gl in
let scheme,elim =
match elim with
| None ->
let sort = Tacticals.New.elimination_sort_of_goal gl in
let _, (elimc,elimt),_ =
guess_elim isrec (* dummy: *) true sort hyp0 gl in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
(* We drop the scheme waiting to know if it is dependent *)
scheme, ElimOver (isrec,hyp0)
| Some e ->
let evd, (elimc,elimt),ind_guess = given_elim hyp0 e gl in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
if Option.is_empty scheme.indarg then error "Cannot find induction type";
- let indsign = compute_scheme_signature scheme hyp0 ind_guess in
+ let indsign = compute_scheme_signature evd scheme hyp0 ind_guess in
let elim = ({elimindex = Some(-1); elimbody = elimc; elimrename = None},elimt) in
scheme, ElimUsing (elim,indsign)
in
@@ -4043,7 +4092,8 @@ let get_elim_signature elim hyp0 gl =
compute_elim_signature (given_elim hyp0 elim gl) hyp0
let is_functional_induction elimc gl =
- let scheme = compute_elim_sig ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in
+ let sigma = Tacmach.New.project gl in
+ let scheme = compute_elim_sig sigma ~elimc (Tacmach.New.pf_unsafe_type_of gl (fst elimc)) in
(* The test is not safe: with non-functional induction on non-standard
induction scheme, this may fail *)
Option.is_empty scheme.indarg
@@ -4058,7 +4108,7 @@ let get_eliminator elim dep s gl =
| ElimOver (isrec,id) ->
let evd, (elimc,elimt),_ as elims = guess_elim isrec dep s id gl in
let _, (l, s) = compute_elim_signature elims id in
- let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (RelDecl.get_type d)))
+ let branchlengthes = List.map (fun d -> assert (RelDecl.is_local_assum d); pi1 (decompose_prod_letin (Tacmach.New.project gl) (RelDecl.get_type d)))
(List.rev s.branches)
in
evd, isrec, ({elimindex = None; elimbody = elimc; elimrename = Some (isrec,Array.of_list branchlengthes)}, elimt), l
@@ -4067,11 +4117,11 @@ let get_eliminator elim dep s gl =
of lid are parameters (first ones), the other are
arguments. Returns the clause obtained. *)
let recolle_clenv i params args elimclause gl =
- let _,arr = destApp elimclause.templval.rebus in
+ let _,arr = destApp elimclause.evd elimclause.templval.rebus in
let lindmv =
Array.map
(fun x ->
- match kind_of_term x with
+ match EConstr.kind elimclause.evd x with
| Meta mv -> mv
| _ -> user_err ~hdr:"elimination_clause"
(str "The type of the elimination clause is not well-formed."))
@@ -4092,7 +4142,7 @@ let recolle_clenv i params args elimclause gl =
(* from_n (Some 0) means that x should be taken "as is" without
trying to unify (which would lead to trying to apply it to
evars if y is a product). *)
- let indclause = Tacmach.New.of_old (fun gl -> mk_clenv_from_n gl (Some 0) (x,y)) gl in
+ let indclause = mk_clenv_from_n gl (Some 0) (x,y) in
let elimclause' = clenv_fchain ~with_univs:false i acc indclause in
elimclause')
(List.rev clauses)
@@ -4103,17 +4153,18 @@ let recolle_clenv i params args elimclause gl =
produce new ones). Then refine with the resulting term with holes.
*)
let induction_tac with_evars params indvars elim =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let ({elimindex=i;elimbody=(elimc,lbindelimc);elimrename=rename},elimt) = elim in
- let i = match i with None -> index_of_ind_arg elimt | Some i -> i in
+ let i = match i with None -> index_of_ind_arg sigma elimt | Some i -> i in
(* elimclause contains this: (elimc ?i ?j ?k...?l) *)
- let elimc = contract_letin_in_lam_header elimc in
+ let elimc = contract_letin_in_lam_header sigma elimc in
let elimc = mkCast (elimc, DEFAULTcast, elimt) in
- let elimclause = pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in
+ let elimclause = Tacmach.New.pf_apply make_clenv_binding gl (elimc,elimt) lbindelimc in
(* elimclause' is built from elimclause by instanciating all args and params. *)
let elimclause' = recolle_clenv i params indvars elimclause gl in
(* one last resolution (useless?) *)
- let resolved = Tacmach.New.of_old (clenv_unique_resolver ~flags:(elim_flags ()) elimclause') gl in
+ let resolved = clenv_unique_resolver ~flags:(elim_flags ()) elimclause' gl in
enforce_prop_bound_names rename (Clenvtac.clenv_refine with_evars resolved)
end }
@@ -4126,9 +4177,9 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
let sigma = Sigma.to_evar_map sigma in
- let concl = Tacmach.New.pf_nf_concl gl in
- let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env in
- let dep_in_concl = Option.cata (fun id -> occur_var env id concl) false hyp0 in
+ let concl = Tacmach.New.pf_concl gl in
+ let statuslists,lhyp0,toclear,deps,avoid,dep_in_hyps = cook_sign hyp0 inhyps indvars env sigma in
+ let dep_in_concl = Option.cata (fun id -> occur_var env sigma id concl) false hyp0 in
let dep = dep_in_hyps || dep_in_concl in
let tmpcl = it_mkNamedProd_or_LetIn concl deps in
let s = Retyping.get_sort_family_of env sigma tmpcl in
@@ -4180,7 +4231,7 @@ let msg_not_right_number_induction_arguments scheme =
must be given, so we help a bit the unifier by making the "pattern"
by hand before calling induction_tac *)
let induction_without_atomization isrec with_evars elim names lid =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let sigma, (indsign,scheme) = get_elim_signature elim (List.hd lid) gl in
let nargs_indarg_farg =
scheme.nargs + (if scheme.farg_in_concl then 1 else 0) in
@@ -4215,8 +4266,8 @@ let induction_without_atomization isrec with_evars elim names lid =
(* assume that no occurrences are selected *)
let clear_unselected_context id inhyps cls =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- if occur_var (Tacmach.New.pf_env gl) id (Tacmach.New.pf_concl gl) &&
+ Proofview.Goal.enter { enter = begin fun gl ->
+ if occur_var (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id (Tacmach.New.pf_concl gl) &&
cls.concl_occs == NoOccurrences
then user_err
(str "Conclusion must be mentioned: it depends on " ++ pr_id id
@@ -4228,7 +4279,7 @@ let clear_unselected_context id inhyps cls =
if Id.List.mem id' inhyps then (* if selected, do not erase *) None
else
(* erase if not selected and dependent on id or selected hyps *)
- let test id = occur_var_in_decl (Tacmach.New.pf_env gl) id d in
+ let test id = occur_var_in_decl (Tacmach.New.pf_env gl) (Tacmach.New.project gl) id d in
if List.exists test (id::inhyps) then Some id' else None in
let ids = List.map_filter to_erase (Proofview.Goal.hyps gl) in
clear ids
@@ -4255,7 +4306,7 @@ let use_bindings env sigma elim must_be_closed (c,lbind) typ =
let rec find_clause typ =
try
let indclause = make_clenv_binding env sigma (c,typ) lbind in
- if must_be_closed && occur_meta (clenv_value indclause) then
+ if must_be_closed && occur_meta indclause.evd (clenv_value indclause) then
error "Need a fully applied argument.";
(* We lose the possibility of coercions in with-bindings *)
let (sigma, c) = pose_all_metas_as_evars env indclause.evd (clenv_value indclause) in
@@ -4273,7 +4324,7 @@ let check_expected_type env sigma (elimc,bl) elimt =
if n == 0 then error "Scheme cannot be applied.";
let sigma,cl = make_evar_clause env sigma ~len:(n - 1) elimt in
let sigma = solve_evar_clause env sigma true cl bl in
- let (_,u,_) = destProd cl.cl_concl in
+ let (_,u,_) = destProd sigma cl.cl_concl in
fun t -> Evarconv.e_cumul env (ref sigma) t u
let check_enough_applied env sigma elim =
@@ -4283,10 +4334,10 @@ let check_enough_applied env sigma elim =
| None ->
(* No eliminator given *)
fun u ->
- let t,_ = decompose_app (whd_all env sigma u) in isInd t
+ let t,_ = decompose_app sigma (whd_all env sigma u) in isInd sigma t
| Some elimc ->
let elimt = Retyping.get_type_of env sigma (fst elimc) in
- let scheme = compute_elim_sig ~elimc elimt in
+ let scheme = compute_elim_sig sigma ~elimc elimt in
match scheme.indref with
| None ->
(* in the absence of information, do not assume it may be
@@ -4305,7 +4356,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Proofview.Goal.env gl in
- let ccl = Proofview.Goal.raw_concl gl in
+ let ccl = Proofview.Goal.concl gl in
let store = Proofview.Goal.extra gl in
let check = check_enough_applied env sigma elim in
let Sigma (c, sigma', p) = use_bindings env sigma elim false (c0,lbind) t0 in
@@ -4336,7 +4387,7 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
end };
if with_evars then Proofview.shelve_unifiable else guard_no_unifiable;
if is_arg_pure_hyp
- then Tacticals.New.tclTRY (clear [destVar c0])
+ then Proofview.tclEVARMAP >>= fun sigma -> Tacticals.New.tclTRY (clear [destVar sigma c0])
else Proofview.tclUNIT ();
if isrec then Proofview.cycle (-1) else Proofview.tclUNIT ()
])
@@ -4360,10 +4411,10 @@ let pose_induction_arg_then isrec with_evars (is_arg_pure_hyp,from_prefix) elim
Sigma (tac, sigma', p +> q)
end }
-let has_generic_occurrences_but_goal cls id env ccl =
+let has_generic_occurrences_but_goal cls id env sigma ccl =
clause_with_generic_context_selection cls &&
(* TODO: whd_evar of goal *)
- (cls.concl_occs != NoOccurrences || not (occur_var env id ccl))
+ (cls.concl_occs != NoOccurrences || not (occur_var env sigma id ccl))
let induction_gen clear_flag isrec with_evars elim
((_pending,(c,lbind)),(eqname,names) as arg) cls =
@@ -4373,14 +4424,15 @@ let induction_gen clear_flag isrec with_evars elim
Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Proofview.Goal.sigma gl in
- let ccl = Proofview.Goal.raw_concl gl in
+ let evd = Sigma.to_evar_map sigma in
+ let ccl = Proofview.Goal.concl gl in
let cls = Option.default allHypsAndConcl cls in
let t = typ_of env sigma c in
let is_arg_pure_hyp =
- isVar c && not (mem_named_context_val (destVar c) (Global.named_context_val ()))
+ isVar evd c && not (mem_named_context_val (destVar evd c) (Global.named_context_val ()))
&& lbind == NoBindings && not with_evars && Option.is_empty eqname
&& clear_flag == None
- && has_generic_occurrences_but_goal cls (destVar c) env ccl in
+ && has_generic_occurrences_but_goal cls (destVar evd c) env evd ccl in
let enough_applied = check_enough_applied env sigma elim t in
if is_arg_pure_hyp && enough_applied then
(* First case: induction on a variable already in an inductive type and
@@ -4388,7 +4440,7 @@ let induction_gen clear_flag isrec with_evars elim
This is a situation where the induction argument is a
clearable variable of the goal w/o occurrence selection
and w/o equality kept: no need to generalize *)
- let id = destVar c in
+ let id = destVar evd c in
Tacticals.New.tclTHEN
(clear_unselected_context id inhyps cls)
(induction_with_atomization_of_ind_arg
@@ -4398,7 +4450,7 @@ let induction_gen clear_flag isrec with_evars elim
declaring the induction argument as a new local variable *)
let id =
(* Type not the right one if partially applied but anyway for internal use*)
- let x = id_of_name_using_hdchar (Global.env()) t Anonymous in
+ let x = id_of_name_using_hdchar (Global.env()) evd t Anonymous in
new_fresh_id [] x gl in
let info_arg = (is_arg_pure_hyp, not enough_applied) in
pose_induction_arg_then
@@ -4423,7 +4475,8 @@ let induction_gen_l isrec with_evars elim names lc =
match l with
| [] -> Proofview.tclUNIT ()
| c::l' ->
- match kind_of_term c with
+ Proofview.tclEVARMAP >>= fun sigma ->
+ match EConstr.kind sigma c with
| Var id when not (mem_named_context_val id (Global.named_context_val ()))
&& not with_evars ->
let _ = newlc:= id::!newlc in
@@ -4432,11 +4485,12 @@ let induction_gen_l isrec with_evars elim names lc =
| _ ->
Proofview.Goal.enter { enter = begin fun gl ->
let type_of = Tacmach.New.pf_unsafe_type_of gl in
+ let sigma = Tacmach.New.project gl in
let x =
- id_of_name_using_hdchar (Global.env()) (type_of c) Anonymous in
+ id_of_name_using_hdchar (Global.env()) sigma (type_of c) Anonymous in
let id = new_fresh_id [] x gl in
- let newl' = List.map (replace_term c (mkVar id)) l' in
+ let newl' = List.map (fun r -> replace_term sigma c (mkVar id) r) l' in
let _ = newlc:=id::!newlc in
Tacticals.New.tclTHEN
(letin_tac None (Name id) c None allHypsAndConcl)
@@ -4458,7 +4512,7 @@ let induction_destruct isrec with_evars (lc,elim) =
match lc with
| [] -> assert false (* ensured by syntax, but if called inside caml? *)
| [c,(eqname,names as allnames),cls] ->
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter { enter = begin fun gl ->
let env = Proofview.Goal.env gl in
let sigma = Tacmach.New.project gl in
match elim with
@@ -4559,9 +4613,9 @@ let simple_destruct = function
*)
let elim_scheme_type elim t =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let clause = Tacmach.New.of_old (fun gl -> mk_clenv_type_of gl elim) gl in
- match kind_of_term (last_arg clause.templval.rebus) with
+ Proofview.Goal.enter { enter = begin fun gl ->
+ let clause = mk_clenv_type_of gl elim in
+ match EConstr.kind clause.evd (last_arg clause.evd clause.templval.rebus) with
| Meta mv ->
let clause' =
(* t is inductive, then CUMUL or CONV is irrelevant *)
@@ -4582,9 +4636,11 @@ let case_type t =
Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let env = Tacmach.New.pf_env gl in
- let (ind,t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in
+ let ((ind, u), t) = reduce_to_atomic_ind env (Sigma.to_evar_map sigma) t in
+ let u = EInstance.kind (Sigma.to_evar_map sigma) u in
let s = Tacticals.New.elimination_sort_of_goal gl in
- let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma ind s in
+ let Sigma (elimc, evd, p) = build_case_analysis_scheme_default env sigma (ind, u) s in
+ let elimc = EConstr.of_constr elimc in
Sigma (elim_scheme_type elimc t, evd, p)
end }
@@ -4598,7 +4654,7 @@ let case_type t =
let (forward_setoid_reflexivity, setoid_reflexivity) = Hook.make ()
let maybe_betadeltaiota_concl allowred gl =
- let concl = Tacmach.New.pf_nf_concl gl in
+ let concl = Tacmach.New.pf_concl gl in
let sigma = Tacmach.New.project gl in
if not allowred then concl
else
@@ -4610,8 +4666,9 @@ let reflexivity_red allowred =
(* PL: usual reflexivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match match_with_equality_type concl with
+ match match_with_equality_type sigma concl with
| None -> Proofview.tclZERO NoEquationFound
| Some _ -> one_constructor 1 NoBindings
end }
@@ -4648,9 +4705,9 @@ let prove_symmetry hdcncl eq_kind =
Tacticals.New.onLastHyp simplest_case;
one_constructor 1 NoBindings ])
-let match_with_equation c =
+let match_with_equation sigma c =
try
- let res = match_with_equation c in
+ let res = match_with_equation sigma c in
Proofview.tclUNIT res
with NoEquationFound ->
Proofview.tclZERO NoEquationFound
@@ -4660,8 +4717,9 @@ let symmetry_red allowred =
(* PL: usual symmetry don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match_with_equation concl >>= fun with_eqn ->
+ match_with_equation sigma concl >>= fun with_eqn ->
match with_eqn with
| Some eq_data,_,_ ->
Tacticals.New.tclTHEN
@@ -4683,16 +4741,18 @@ let (forward_setoid_symmetry_in, setoid_symmetry_in) = Hook.make ()
let symmetry_in id =
Proofview.Goal.enter { enter = begin fun gl ->
+ let sigma = Tacmach.New.project gl in
let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in
- let sign,t = decompose_prod_assum ctype in
+ let sign,t = decompose_prod_assum sigma ctype in
Proofview.tclORELSE
begin
- match_with_equation t >>= fun (_,hdcncl,eq) ->
- let symccl = match eq with
+ match_with_equation sigma t >>= fun (_,hdcncl,eq) ->
+ let symccl =
+ match eq with
| MonomorphicLeibnizEq (c1,c2) -> mkApp (hdcncl, [| c2; c1 |])
| PolymorphicLeibnizEq (typ,c1,c2) -> mkApp (hdcncl, [| typ; c2; c1 |])
| HeterogenousEq (t1,c1,t2,c2) -> mkApp (hdcncl, [| t2; c2; t1; c1 |]) in
- Tacticals.New.tclTHENS (cut (it_mkProd_or_LetIn symccl sign))
+ Tacticals.New.tclTHENS (cut (EConstr.it_mkProd_or_LetIn symccl sign))
[ intro_replacing id;
Tacticals.New.tclTHENLIST [ intros; symmetry; apply (mkVar id); assumption ] ]
end
@@ -4752,8 +4812,9 @@ let transitivity_red allowred t =
(* PL: usual transitivity don't perform any reduction when searching
for an equality, but we may need to do some when called back from
inside setoid_reflexivity (see Optimize cases in setoid_replace.ml). *)
+ let sigma = Tacmach.New.project gl in
let concl = maybe_betadeltaiota_concl allowred gl in
- match_with_equation concl >>= fun with_eqn ->
+ match_with_equation sigma concl >>= fun with_eqn ->
match with_eqn with
| Some eq_data,_,_ ->
Tacticals.New.tclTHEN
@@ -4787,6 +4848,12 @@ let intros_transitivity n = Tacticals.New.tclTHEN intros (transitivity_gen n)
(** d1 is the section variable in the global context, d2 in the goal context *)
let interpretable_as_section_decl evd d1 d2 =
let open Context.Named.Declaration in
+ let e_eq_constr_univs sigma c1 c2 = match eq_constr_universes !sigma c1 c2 with
+ | None -> false
+ | Some cstr ->
+ try ignore (Evd.add_universe_constraints !sigma cstr); true
+ with UniversesDiffer -> false
+ in
match d2, d1 with
| LocalDef _, LocalAssum _ -> false
| LocalDef (_,b1,t1), LocalDef (_,b2,t2) ->
@@ -4804,6 +4871,8 @@ let rec decompose len c t accu =
| _ -> assert false
let rec shrink ctx sign c t accu =
+ let open Term in
+ let open CVars in
match ctx, sign with
| [], [] -> (c, t, accu)
| p :: ctx, decl :: sign ->
@@ -4842,7 +4911,7 @@ let abstract_subproof id gk tac =
let open Tacticals.New in
let open Tacmach.New in
let open Proofview.Notations in
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
let current_sign = Global.named_context_val ()
and global_sign = Proofview.Goal.hyps gl in
@@ -4870,6 +4939,7 @@ let abstract_subproof id gk tac =
let ctx = Evd.universe_context_set evd in
evd, ctx, nf concl
in
+ let concl = EConstr.of_constr concl in
let solve_tac = tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac) in
let ectx = Evd.evar_universe_context evd in
let (const, safe, ectx) =
@@ -4884,8 +4954,9 @@ let abstract_subproof id gk tac =
in
let const, args =
if !shrink_abstract then shrink_entry sign const
- else (const, List.rev (Context.Named.to_instance sign))
+ else (const, List.rev (Context.Named.to_instance Constr.mkVar sign))
in
+ let args = List.map EConstr.of_constr args in
let cd = Entries.DefinitionEntry const in
let decl = (cd, IsProof Lemma) in
let cst () =
@@ -4897,6 +4968,7 @@ let abstract_subproof id gk tac =
let cst = Impargs.with_implicit_protection cst () in
(* let evd, lem = Evd.fresh_global (Global.env ()) evd (ConstRef cst) in *)
let lem, ctx = Universes.unsafe_constr_of_global (ConstRef cst) in
+ let lem = EConstr.of_constr lem in
let evd = Evd.set_universe_context evd ectx in
let open Safe_typing in
let eff = private_con_of_con (Global.safe_env ()) cst in
@@ -4928,7 +5000,7 @@ let tclABSTRACT name_op tac =
abstract_subproof s gk tac
let unify ?(state=full_transparent_state) x y =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.s_enter { s_enter = begin fun gl ->
let sigma = Proofview.Goal.sigma gl in
try
let core_flags =
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 354ac50b4..ba4a9706d 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -9,6 +9,7 @@
open Loc
open Names
open Term
+open EConstr
open Environ
open Proof_type
open Evd
@@ -28,15 +29,15 @@ open Locus
(** {6 General functions. } *)
-val is_quantified_hypothesis : Id.t -> ([`NF],'b) Proofview.Goal.t -> bool
+val is_quantified_hypothesis : Id.t -> ('a, 'r) Proofview.Goal.t -> bool
(** {6 Primitive tactics. } *)
val introduction : ?check:bool -> Id.t -> unit Proofview.tactic
val convert_concl : ?check:bool -> types -> cast_kind -> unit Proofview.tactic
-val convert_hyp : ?check:bool -> Context.Named.Declaration.t -> unit Proofview.tactic
+val convert_hyp : ?check:bool -> named_declaration -> unit Proofview.tactic
val convert_concl_no_check : types -> cast_kind -> unit Proofview.tactic
-val convert_hyp_no_check : Context.Named.Declaration.t -> unit Proofview.tactic
+val convert_hyp_no_check : named_declaration -> unit Proofview.tactic
val mutual_fix :
Id.t -> int -> (Id.t * int * constr) list -> int -> unit Proofview.tactic
val fix : Id.t option -> int -> unit Proofview.tactic
@@ -50,7 +51,7 @@ val convert_leq : constr -> constr -> unit Proofview.tactic
val fresh_id_in_env : Id.t list -> Id.t -> env -> Id.t
val fresh_id : Id.t list -> Id.t -> goal sigma -> Id.t
-val find_intro_names : Context.Rel.t -> goal sigma -> Id.t list
+val find_intro_names : rel_context -> goal sigma -> Id.t list
val intro : unit Proofview.tactic
val introf : unit Proofview.tactic
@@ -74,7 +75,7 @@ val intros : unit Proofview.tactic
(** [depth_of_quantified_hypothesis b h g] returns the index of [h] in
the conclusion of goal [g], up to head-reduction if [b] is [true] *)
val depth_of_quantified_hypothesis :
- bool -> quantified_hypothesis -> ([`NF],'b) Proofview.Goal.t -> int
+ bool -> quantified_hypothesis -> ('a, 'r) Proofview.Goal.t -> int
val intros_until : quantified_hypothesis -> unit Proofview.tactic
@@ -184,7 +185,7 @@ val revert : Id.t list -> unit Proofview.tactic
(** {6 Resolution tactics. } *)
val apply_type : constr -> constr list -> unit Proofview.tactic
-val bring_hyps : Context.Named.t -> unit Proofview.tactic
+val bring_hyps : named_context -> unit Proofview.tactic
val apply : constr -> unit Proofview.tactic
val eapply : constr -> unit Proofview.tactic
@@ -243,15 +244,15 @@ type elim_scheme = {
elimc: constr with_bindings option;
elimt: types;
indref: global_reference option;
- params: Context.Rel.t; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
+ params: rel_context; (** (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *)
nparams: int; (** number of parameters *)
- predicates: Context.Rel.t; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
+ predicates: rel_context; (** (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *)
npredicates: int; (** Number of predicates *)
- branches: Context.Rel.t; (** branchr,...,branch1 *)
+ branches: rel_context; (** branchr,...,branch1 *)
nbranches: int; (** Number of branches *)
- args: Context.Rel.t; (** (xni, Ti_ni) ... (x1, Ti_1) *)
+ args: rel_context; (** (xni, Ti_ni) ... (x1, Ti_1) *)
nargs: int; (** number of arguments *)
- indarg: Context.Rel.Declaration.t option; (** Some (H,I prm1..prmp x1...xni)
+ indarg: rel_declaration option; (** Some (H,I prm1..prmp x1...xni)
if HI is in premisses, None otherwise *)
concl: types; (** Qi x1...xni HI (f...), HI and (f...)
are optional and mutually exclusive *)
@@ -259,7 +260,7 @@ type elim_scheme = {
farg_in_concl: bool; (** true if (f...) appears at the end of conclusion *)
}
-val compute_elim_sig : ?elimc: constr with_bindings -> types -> elim_scheme
+val compute_elim_sig : evar_map -> ?elimc:constr with_bindings -> types -> elim_scheme
(** elim principle with the index of its inductive arg *)
type eliminator = {
diff --git a/tactics/term_dnet.ml b/tactics/term_dnet.ml
index e4b45489d..2c863c42a 100644
--- a/tactics/term_dnet.ml
+++ b/tactics/term_dnet.ml
@@ -344,18 +344,19 @@ struct
) (pr_dconstr pr_term_pattern) p*)
let search_pat cpat dpat dn =
- let whole_c = cpat in
+ let whole_c = EConstr.of_constr cpat in
(* if we are at the root, add an empty context *)
let dpat = under_prod (empty_ctx dpat) in
TDnet.Idset.fold
(fun id acc ->
let c_id = Opt.reduce (Ident.constr_of id) in
+ let c_id = EConstr.of_constr c_id in
let (ctx,wc) =
- try Termops.align_prod_letin whole_c c_id
+ try Termops.align_prod_letin Evd.empty whole_c c_id (** FIXME *)
with Invalid_argument _ -> [],c_id in
let wc,whole_c = if Opt.direction then whole_c,wc else wc,whole_c in
try
- let _ = Termops.filtering ctx Reduction.CUMUL wc whole_c in
+ let _ = Termops.filtering Evd.empty ctx Reduction.CUMUL wc whole_c in
id :: acc
with Termops.CannotFilter -> (* msgnl(str"recon "++Termops.print_constr_env (Global.env()) wc); *) acc
) (TDnet.find_match dpat dn) []