summaryrefslogtreecommitdiff
path: root/tactics/equality.ml
diff options
context:
space:
mode:
Diffstat (limited to 'tactics/equality.ml')
-rw-r--r--tactics/equality.ml714
1 files changed, 365 insertions, 349 deletions
diff --git a/tactics/equality.ml b/tactics/equality.ml
index bb3cbad9..5821717d 100644
--- a/tactics/equality.ml
+++ b/tactics/equality.ml
@@ -1,23 +1,26 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
+module CVars = Vars
open Pp
open CErrors
open Util
open Names
open Nameops
open Term
-open Vars
open Termops
+open EConstr
+open Vars
open Namegen
open Inductive
open Inductiveops
-open Environ
open Libnames
open Globnames
open Reductionops
@@ -26,7 +29,6 @@ open Retyping
open Tacmach.New
open Logic
open Hipattern
-open Tacexpr
open Tacticals.New
open Tactics
open Tacred
@@ -40,53 +42,35 @@ open Eqschemes
open Locus
open Locusops
open Misctypes
-open Sigma.Notations
open Proofview.Notations
open Unification
open Context.Named.Declaration
-(* Options *)
+module NamedDecl = Context.Named.Declaration
-let discriminate_introduction = ref true
+(* Options *)
-let discr_do_intro () =
- !discriminate_introduction && Flags.version_strictly_greater Flags.V8_2
+type inj_flags = {
+ keep_proof_equalities : bool;
+ injection_in_context : bool;
+ injection_pattern_l2r_order : bool;
+ }
open Goptions
-let _ =
- declare_bool_option
- { optsync = true;
- optdepr = false;
- optname = "automatic introduction of hypotheses by discriminate";
- optkey = ["Discriminate";"Introduction"];
- optread = (fun () -> !discriminate_introduction);
- optwrite = (:=) discriminate_introduction }
-
-let injection_pattern_l2r_order = ref true
-let use_injection_pattern_l2r_order () =
- !injection_pattern_l2r_order
- && Flags.version_strictly_greater Flags.V8_4
-
-let _ =
- declare_bool_option
- { optsync = true;
- optdepr = false;
- optname = "injection left-to-right pattern order and clear by default when with introduction pattern";
- optkey = ["Injection";"L2R";"Pattern";"Order"];
- optread = (fun () -> !injection_pattern_l2r_order) ;
- optwrite = (fun b -> injection_pattern_l2r_order := b) }
+let use_injection_pattern_l2r_order = function
+ | None -> true
+ | Some flags -> flags.injection_pattern_l2r_order
let injection_in_context = ref false
-let use_injection_in_context () =
- !injection_in_context
- && Flags.version_strictly_greater Flags.V8_5
+let use_injection_in_context = function
+ | None -> !injection_in_context
+ | Some flags -> flags.injection_in_context
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "injection in context";
optkey = ["Structural";"Injection"];
optread = (fun () -> !injection_in_context) ;
@@ -94,9 +78,6 @@ let _ =
(* Rewriting tactics *)
-let tclNOTSAMEGOAL tac =
- Proofview.V82.tactic (Tacticals.tclNOTSAMEGOAL (Proofview.V82.of_tactic tac))
-
type dep_proof_flag = bool (* true = support rewriting dependent proofs *)
type freeze_evars_flag = bool (* true = don't instantiate existing evars *)
@@ -143,7 +124,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
@@ -164,10 +145,11 @@ 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 () = if arglen < 2 then user_err Pp.(str "The term provided is not an applied relation.") in
let c1 = args.(arglen - 2) in
let c2 = args.(arglen - 1) in
let try_occ (evd', c') =
@@ -257,12 +239,31 @@ let rewrite_keyed_unif_flags = {
}
let rewrite_elim with_evars frzevars cls c e =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let flags = if Unification.is_keyed_unification ()
then rewrite_keyed_unif_flags else rewrite_conv_closed_unif_flags in
let flags = make_flags frzevars (Tacmach.New.project gl) flags c in
general_elim_clause with_evars flags cls c e
- end }
+ end
+
+let tclNOTSAMEGOAL tac =
+ let goal gl = Proofview.Goal.goal gl in
+ Proofview.Goal.nf_enter begin fun gl ->
+ let sigma = project gl in
+ let ev = goal gl in
+ tac >>= fun () ->
+ Proofview.Goal.goals >>= fun gls ->
+ let check accu gl' =
+ gl' >>= fun gl' ->
+ let accu = accu || Goal.V82.same_goal sigma ev (project gl') (goal gl') in
+ Proofview.tclUNIT accu
+ in
+ Proofview.Monad.List.fold_left check false gls >>= fun has_same ->
+ if has_same then
+ tclZEROMSG (str"Tactic generated a subgoal identical to the original goal.")
+ else
+ Proofview.tclUNIT ()
+ end
(* Ad hoc asymmetric general_elim_clause *)
let general_elim_clause with_evars frzevars cls rew elim =
@@ -297,19 +298,19 @@ let general_elim_clause with_evars frzevars tac cls c t l l2r elim =
(general_elim_clause with_evars frzevars cls c elim))
tac
in
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let instantiate_lemma concl =
if not all then instantiate_lemma gl c t l l2r concl
else instantiate_lemma_all frzevars gl c t l l2r concl
in
let typ = match cls with
- | None -> pf_nf_concl gl
- | Some id -> pf_get_hyp_typ id (Proofview.Goal.assume gl)
+ | None -> pf_concl gl
+ | Some id -> pf_get_hyp_typ id gl
in
let cs = instantiate_lemma typ in
if firstonly then tclFIRST (List.map try_clause cs)
else tclMAP try_clause cs
- end }
+ end
(* The next function decides in particular whether to try a regular
rewrite or a generalized rewrite.
@@ -322,27 +323,30 @@ let (forward_general_setoid_rewrite_clause, general_setoid_rewrite_clause) = Hoo
(* Do we have a JMeq instance on twice the same domains ? *)
-let jmeq_same_dom gl = function
+let jmeq_same_dom env sigma = 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
- | _, [dom1; _; dom2;_] -> is_conv env (Tacmach.New.project gl) dom1 dom2
+ let rels, t = decompose_prod_assum sigma t in
+ let env = push_rel_context rels env in
+ match decompose_app sigma t with
+ | _, [dom1; _; dom2;_] -> is_conv env sigma dom1 dom2
| _ -> false
(* find_elim determines which elimination principle is necessary to
eliminate lbeq on sort_of_gl. *)
-let find_elim hdcncl lft2rgt dep cls ot gl =
+let find_elim hdcncl lft2rgt dep cls ot =
+ Proofview.Goal.enter_one begin fun 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
+ let env = Proofview.Goal.env gl in
if (is_global Coqlib.glob_eq hdcncl ||
(is_global Coqlib.glob_jmeq hdcncl &&
- jmeq_same_dom gl ot)) && not dep
- || Flags.version_less_or_equal Flags.V8_2
+ jmeq_same_dom env sigma ot)) && not dep
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)
@@ -351,16 +355,16 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
| Some true, None
| Some false, Some _ ->
let c1 = destConstRef pr1 in
- let mp,dp,l = repr_con (constant_of_kn (canonical_con c1)) in
+ let mp,dp,l = Constant.repr3 (Constant.make1 (Constant.canonical c1)) in
let l' = Label.of_id (add_suffix (Label.to_id l) "_r") in
- let c1' = Global.constant_of_delta_kn (make_kn mp dp l') in
+ let c1' = Global.constant_of_delta_kn (KerName.make mp dp l') in
begin
try
let _ = Global.lookup_constant c1' in
c1'
with Not_found ->
- errorlabstrm "Equality.find_elim"
- (str "Cannot find rewrite principle " ++ pr_label l' ++ str ".")
+ user_err ~hdr:"Equality.find_elim"
+ (str "Cannot find rewrite principle " ++ Label.print l' ++ str ".")
end
| _ -> destConstRef pr1
end
@@ -369,8 +373,7 @@ let find_elim hdcncl lft2rgt dep cls ot gl =
Logic.eq or Jmeq just before *)
assert false
in
- let Sigma (elim, sigma, p) = Sigma.fresh_global (Global.env ()) (Proofview.Goal.sigma gl) (ConstRef c) in
- Sigma ((elim, Safe_typing.empty_private_constants), sigma, p)
+ pf_constr_of_global (ConstRef c)
else
let scheme_name = match dep, lft2rgt, inccl with
(* Non dependent case *)
@@ -384,36 +387,31 @@ 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
- Sigma ((elim, eff), sigma, p)
+ Proofview.tclEFFECTS eff <*>
+ pf_constr_of_global (ConstRef c)
| _ -> assert false
+ end
let type_of_clause cls gl = match cls with
| None -> Proofview.Goal.concl gl
| 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 ->
- let evd = Sigma.to_evar_map (Proofview.Goal.sigma gl) in
- let isatomic = isProd (whd_zeta evd hdcncl) in
+ Proofview.Goal.enter begin fun gl ->
+ let evd = Proofview.Goal.sigma gl 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 Sigma ((elim, effs), sigma, p) = find_elim hdcncl lft2rgt dep cls (Some t) gl in
- let tac =
- Proofview.tclEFFECTS effs <*>
+ let dep = dep_proof_ok && dep_fun evd c type_of_cls in
+ find_elim hdcncl lft2rgt dep cls (Some t) >>= fun elim ->
general_elim_clause with_evars frzevars tac cls c t l
(match lft2rgt with None -> false | Some b -> b)
{elimindex = None; elimbody = (elim,NoBindings); elimrename = None}
- in
- Sigma (tac, sigma, p)
- end }
+ end
let adjust_rewriting_direction args lft2rgt =
match args with
@@ -421,7 +419,7 @@ let adjust_rewriting_direction args lft2rgt =
(* equality to a constant, like in eq_true *)
(* more natural to see -> as the rewriting to the constant *)
if not lft2rgt then
- error "Rewriting non-symmetric equality not allowed from right-to-left.";
+ user_err Pp.(str "Rewriting non-symmetric equality not allowed from right-to-left.");
None
| _ ->
(* other equality *)
@@ -436,12 +434,12 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
if occs != AllOccurrences then (
rewrite_side_tac (Hook.get forward_general_setoid_rewrite_clause cls lft2rgt occs (c,l) ~new_goals:[]) tac)
else
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
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)
@@ -454,9 +452,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
@@ -464,7 +463,7 @@ let general_rewrite_ebindings_clause cls lft2rgt occs frzevars dep_proof_ok ?tac
| None -> Proofview.tclZERO ~info e
(* error "The provided term does not end with an equality or a declared rewrite relation." *)
end
- end }
+ end
let general_rewrite_ebindings =
general_rewrite_ebindings_clause None
@@ -515,46 +514,46 @@ let general_rewrite_clause l2r with_evars ?tac c cl =
let rec do_hyps_atleastonce = function
| [] -> tclZEROMSG (Pp.str"Nothing to rewrite.")
| id :: l ->
- tclIFTHENTRYELSEMUST
+ tclIFTHENFIRSTTRYELSEMUST
(general_rewrite_ebindings_in l2r AllOccurrences false true ?tac id c with_evars)
(do_hyps_atleastonce l)
in
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 (Proofview.Goal.env gl) (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
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
do_hyps_atleastonce (ids gl)
- end }
+ end
in
if cl.concl_occs == NoOccurrences then do_hyps else
- tclIFTHENTRYELSEMUST
+ tclIFTHENFIRSTTRYELSEMUST
(general_rewrite_ebindings l2r (occs_of cl.concl_occs) false true ?tac c with_evars)
do_hyps
let apply_special_clear_request clear_flag f =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
try
- let ((c, bl), sigma) = run_delayed env sigma f in
+ let (sigma, (c, bl)) = f env sigma in
apply_clear_request clear_flag (use_clear_hyp_by_default ()) c
with
e when catchable_exception e -> tclIDTAC
- end }
+ end
let general_multi_rewrite with_evars l cl tac =
let do1 l2r f =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
- let (c, sigma) = run_delayed env sigma f in
+ let (sigma, c) = f env sigma in
tclWITHHOLES with_evars
(general_rewrite_clause l2r with_evars ?tac c cl) sigma
- end }
+ end
in
let rec doN l2r c = function
| Precisely n when n <= 0 -> Proofview.tclUNIT ()
@@ -617,7 +616,7 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt =
| None -> Proofview.tclUNIT ()
| Some tac -> tclCOMPLETE tac
in
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let get_type_of = pf_apply get_type_of gl in
let t1 = get_type_of c1
and t2 = get_type_of c2 in
@@ -633,8 +632,8 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt =
| Some evd ->
let e = build_coq_eq () in
let sym = build_coq_eq_sym () in
- Tacticals.New.pf_constr_of_global sym (fun sym ->
- Tacticals.New.pf_constr_of_global e (fun e ->
+ Tacticals.New.pf_constr_of_global sym >>= fun sym ->
+ Tacticals.New.pf_constr_of_global e >>= fun e ->
let eq = applist (e, [t1;c1;c2]) in
tclTHENLAST
(replace_core clause l2r eq)
@@ -642,8 +641,8 @@ let replace_using_leibniz clause c1 c2 l2r unsafe try_prove_eq_opt =
[assumption;
tclTHEN (apply sym) assumption;
try_prove_eq
- ])))
- end }
+ ])
+ end
let replace c1 c2 =
replace_using_leibniz onConcl c2 c1 false false None
@@ -705,26 +704,31 @@ let keep_proof_equalities_for_injection = ref false
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "injection on prop arguments";
optkey = ["Keep";"Proof";"Equalities"];
optread = (fun () -> !keep_proof_equalities_for_injection) ;
optwrite = (fun b -> keep_proof_equalities_for_injection := b) }
+let keep_proof_equalities = function
+ | None -> !keep_proof_equalities_for_injection
+ | Some flags -> flags.keep_proof_equalities
-let find_positions env sigma t1 t2 =
+(* [keep_proofs] is relevant for types in Prop with elimination in Type *)
+(* In particular, it is relevant for injection but not for discriminate *)
+
+let find_positions env sigma ~keep_proofs ~no_discr t1 t2 =
let project env sorts posn t1 t2 =
let ty1 = get_type_of env sigma t1 in
- let s = get_sort_family_of env sigma ty1 in
+ let s = get_sort_family_of ~truncation_style:true env sigma ty1 in
if Sorts.List.mem s sorts
then [(List.rev posn,t1,t2)] else []
in
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
- | Construct (sp1,_), Construct (sp2,_)
+ 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)
->
let sorts' =
@@ -733,13 +737,18 @@ let find_positions env sigma t1 t2 =
(* both sides are fully applied constructors, so either we descend,
or we can discriminate here. *)
if eq_constructor sp1 sp2 then
- let nrealargs = constructor_nrealargs_env env sp1 in
- let rargs1 = List.lastn nrealargs args1 in
- let rargs2 = List.lastn nrealargs args2 in
+ let nparams = inductive_nparams_env env ind1 in
+ 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 = CVars.adjust_rel_to_rel_context ctxt (i+1) - 1 in
List.flatten
- (List.map2_i (fun i -> findrec sorts' ((sp1,i)::posn))
+ (List.map2_i (fun i -> findrec sorts' ((sp1,adjust i)::posn))
0 rargs1 rargs2)
- else if Sorts.List.mem InType sorts'
+ else if Sorts.List.mem InType sorts' && not no_discr
then (* see build_discriminator *)
raise (DiscrFound (List.rev posn,sp1,sp2))
else
@@ -755,21 +764,24 @@ let find_positions env sigma t1 t2 =
project env sorts posn t1_0 t2_0
in
try
- let sorts = if !keep_proof_equalities_for_injection then [InSet;InType;InProp]
- else [InSet;InType]
- in
+ let sorts = if keep_proofs then [InSet;InType;InProp] else [InSet;InType] in
Inr (findrec sorts [] t1 t2)
with DiscrFound (path,c1,c2) ->
Inl (path,c1,c2)
+let use_keep_proofs = function
+ | None -> !keep_proof_equalities_for_injection
+ | Some b -> b
+
let discriminable env sigma t1 t2 =
- match find_positions env sigma t1 t2 with
+ match find_positions env sigma ~keep_proofs:false ~no_discr:false t1 t2 with
| Inl _ -> true
| _ -> false
-let injectable env sigma t1 t2 =
- match find_positions env sigma t1 t2 with
- | Inl _ | Inr [] | Inr [([],_,_)] -> false
+let injectable env sigma ~keep_proofs t1 t2 =
+ match find_positions env sigma ~keep_proofs:(use_keep_proofs keep_proofs) ~no_discr:true t1 t2 with
+ | Inl _ -> assert false
+ | Inr [] | Inr [([],_,_)] -> false
| Inr _ -> true
@@ -840,27 +852,30 @@ let descend_then env sigma head dirn =
let IndType (indf,_) =
try find_rectype env sigma (get_type_of env sigma head)
with Not_found ->
- error "Cannot project on an inductive type derived from a dependency." in
+ user_err Pp.(str "Cannot project on an inductive type derived from a dependency.")
+ in
let indp,_ = (dest_ind_family indf) in
let ind, _ = check_privacy env indp in
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
+ (fun sigma dirnval (dfltval,resty) ->
+ 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:
@@ -888,35 +903,39 @@ let build_selector env sigma dirn c ind special default =
on (c bool true) = (c bool false)
CP : changed assert false in a more informative error
*)
- errorlabstrm "Equality.construct_discriminator"
+ user_err ~hdr:"Equality.construct_discriminator"
(str "Cannot discriminate on inductive constructors with \
dependent types.") in
let (indp,_) = dest_ind_family indf in
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 ans = Inductiveops.make_case_or_project env sigma indf ci p c (Array.of_list brl) in
+ ans
-let rec build_discriminator env sigma dirn c = function
+let build_coq_False () = pf_constr_of_global (build_coq_False ())
+let build_coq_True () = pf_constr_of_global (build_coq_True ())
+let build_coq_I () = pf_constr_of_global (build_coq_I ())
+
+let rec build_discriminator env sigma true_0 false_0 dirn c = function
| [] ->
let ind = get_type_of env sigma c in
- let true_0,false_0 =
- build_coq_True(),build_coq_False() in
build_selector env sigma dirn c ind true_0 false_0
| ((sp,cnum),argnum)::l ->
let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in
let newc = mkRel(cnum_nlams-argnum) in
- let subval = build_discriminator cnum_env sigma dirn newc l in
- kont subval (build_coq_False (),mkSort (Prop Null))
+ let subval = build_discriminator cnum_env sigma true_0 false_0 dirn newc l in
+ kont sigma subval (false_0,mkSort (Prop Null))
(* Note: discrimination could be more clever: if some elimination is
not allowed because of a large impredicative constructor in the
@@ -930,15 +949,15 @@ let rec build_discriminator env sigma dirn c = function
*)
let gen_absurdity id =
- Proofview.Goal.enter { enter = begin fun gl ->
- 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
+ Proofview.Goal.enter begin fun gl ->
+ let sigma = project gl in
+ let hyp_typ = pf_get_hyp_typ id gl in
+ if is_empty_type sigma hyp_typ
then
simplest_elim (mkVar id)
else
tclZEROMSG (str "Not the negation of an equality.")
- end }
+ end
(* Precondition: eq is leibniz equality
@@ -958,13 +977,15 @@ let ind_scheme_of_eq lbeq =
ConstRef c, eff
-let discrimination_pf env sigma e (t,t1,t2) discriminator lbeq =
- let i = build_coq_I () in
- 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
- sigma, (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term),
- eff
+let discrimination_pf e (t,t1,t2) discriminator lbeq =
+ build_coq_I () >>= fun i ->
+ build_coq_False () >>= fun absurd_term ->
+ let eq_elim, eff = ind_scheme_of_eq lbeq in
+ Proofview.tclEFFECTS eff <*>
+ pf_constr_of_global eq_elim >>= fun eq_elim ->
+ Proofview.tclUNIT
+ (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term)
+
let eq_baseid = Id.of_string "e"
@@ -972,39 +993,44 @@ 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
- | _ -> errorlabstrm "" (str "Ill-formed clause applicator.")) in
+ | _ -> user_err (str "Ill-formed clause applicator.")) in
clenv_fchain ~with_univs:false argmv f_clause clause
let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn =
- let e = next_ident_away eq_baseid (ids_of_context env) in
+ build_coq_True () >>= fun true_0 ->
+ build_coq_False () >>= fun false_0 ->
+ let e = next_ident_away eq_baseid (vars_of_env env) in
let e_env = push_named (Context.Named.Declaration.LocalAssum (e,t)) env in
let discriminator =
- build_discriminator e_env sigma dirn (mkVar e) cpath in
- let sigma,(pf, absurd_term), eff =
- discrimination_pf env sigma e (t,t1,t2) discriminator lbeq in
- let pf_ty = mkArrow eqn absurd_term in
- let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in
- let pf = Clenvtac.clenv_value_cast_meta absurd_clause in
- Proofview.Unsafe.tclEVARS sigma <*>
- Proofview.tclEFFECTS eff <*>
- tclTHENS (assert_after Anonymous absurd_term)
- [onLastHypId gen_absurdity; (Proofview.V82.tactic (Tacmach.refine pf))]
+ try
+ Proofview.tclUNIT
+ (build_discriminator e_env sigma true_0 false_0 dirn (mkVar e) cpath)
+ with
+ UserError _ as ex -> Proofview.tclZERO ex
+ in
+ discriminator >>= fun discriminator ->
+ discrimination_pf e (t,t1,t2) discriminator lbeq >>= fun (pf, absurd_term) ->
+ let pf_ty = mkArrow eqn absurd_term in
+ let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in
+ let pf = Clenvtac.clenv_value_cast_meta absurd_clause in
+ tclTHENS (assert_after Anonymous absurd_term)
+ [onLastHypId gen_absurdity; (Proofview.V82.tactic (Tacmach.refine pf))]
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 begin fun gl ->
let env = Proofview.Goal.env gl in
- match find_positions env sigma t1 t2 with
+ match find_positions env sigma ~keep_proofs:false ~no_discr:false t1 t2 with
| Inr _ ->
tclZEROMSG (str"Not a discriminable equality.")
| Inl (cpath, (_,dirn), _) ->
discr_positions env sigma u eq_clause cpath dirn
- end }
+ end
let onEquality with_evars tac (c,lbindc) =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.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
@@ -1016,21 +1042,21 @@ let onEquality with_evars tac (c,lbindc) =
tclTHEN
(Proofview.Unsafe.tclEVARS eq_clause'.evd)
(tac (eq,eqn,eq_args) eq_clause')
- end }
+ end
let onNegatedEquality with_evars tac =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.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)))
| _ ->
tclZEROMSG (str "Not a negated primitive equality.")
- end }
+ end
let discrSimpleClause with_evars = function
| None -> onNegatedEquality with_evars discrEq
@@ -1041,19 +1067,13 @@ let discr with_evars = onEquality with_evars discrEq
let discrClause with_evars = onClause (discrSimpleClause with_evars)
let discrEverywhere with_evars =
-(*
- tclORELSE
-*)
- (if discr_do_intro () then
- (tclTHEN
- (tclREPEAT introf)
- (tryAllHyps
+ tclTHEN (Proofview.tclUNIT ())
+ (* Delay the interpretation of side-effect *)
+ (tclTHEN
+ (tclREPEAT introf)
+ (tryAllHyps
(fun id -> tclCOMPLETE (discr with_evars (mkVar id,NoBindings)))))
- else (* <= 8.2 compat *)
- tryAllHypsAndConcl (discrSimpleClause with_evars))
-(* (fun gls ->
- errorlabstrm "DiscrEverywhere" (str"No discriminable equalities."))
-*)
+
let discr_tac with_evars = function
| None -> discrEverywhere with_evars
| Some c -> onInductionArg (fun clear_flag -> discr with_evars) c
@@ -1078,7 +1098,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
@@ -1088,6 +1108,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]))
@@ -1100,9 +1122,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
@@ -1167,25 +1189,24 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt =
evdref := Evarconv.solve_unif_constraints_with_heuristics env !evdref in
dflt
with Evarconv.UnableToUnify _ ->
- error "Cannot solve a unification problem."
+ user_err Pp.(str "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)
- | _ -> anomaly ~label:"sig_clausal_form" (Pp.str "should be a sigma type") in
+ | (_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."
+ user_err Pp.(str "Cannot solve a unification problem.")
| None ->
(* This at least happens if what has been detected as a
dependency is not one; use an evasive error message;
@@ -1193,7 +1214,7 @@ let sig_clausal_form env sigma sort_of_ty siglen ty dflt =
tried in the first place in make_iterated_tuple instead
of approximatively computing the free rels; then
unsolved evars would mean not binding rel *)
- error "Cannot solve a unification problem."
+ user_err Pp.(str "Cannot solve a unification problem.")
in
let scf = sigrec_clausal_form siglen ty in
!evdref, Evarutil.nf_evar !evdref scf
@@ -1263,7 +1284,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)
@@ -1275,7 +1296,8 @@ let rec build_injrec env sigma dflt c = function
let (cnum_nlams,cnum_env,kont) = descend_then env sigma c cnum in
let newc = mkRel(cnum_nlams-argnum) in
let sigma, (subval,tuplety,dfltval) = build_injrec cnum_env sigma dflt newc l in
- sigma, (kont subval (dfltval,tuplety), tuplety,dfltval)
+ let res = kont sigma subval (dfltval,tuplety) in
+ sigma, (res, tuplety,dfltval)
with
UserError _ -> failwith "caught"
@@ -1283,62 +1305,51 @@ 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 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 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 = Coqlib.coq_reference "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"]
+ "inj_pair2_eq_dec" 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 ->
+ Tacticals.New.pf_constr_of_global Coqlib.glob_eq >>= fun ceq ->
tclTHENS (cut (mkApp (ceq,new_eq_args)))
- [clear [destVar hyp];
+ [clear [destVar sigma hyp];
+ Tacticals.New.pf_constr_of_global inj2 >>= fun inj2 ->
Proofview.V82.tactic (Tacmach.refine
(mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3);hyp|])))
])]
with Exit ->
Proofview.tclUNIT ()
- end }
+ end
(* Given t1=t2 Inj calculates the whd normal forms of t1 and t2 and it
expands then only when the whdnf has a constructor of an inductive type
@@ -1346,13 +1357,13 @@ 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
let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
- let e = next_ident_away eq_baseid (ids_of_context env) in
+ let e = next_ident_away eq_baseid (vars_of_env env) in
let e_env = push_named (LocalAssum (e,t)) env in
let evdref = ref sigma in
let filter (cpath, t1', t2') =
@@ -1361,6 +1372,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
@@ -1383,46 +1395,48 @@ let inject_at_positions env sigma l2r (eq,_,(t,t1,t2)) eq_clause posns tac =
(if l2r then List.rev injectors else injectors)))
(tac (List.length injectors)))
-let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause =
+let injEqThen keep_proofs tac l2r (eq,_,(t,t1,t2) as u) eq_clause =
let sigma = eq_clause.evd in
let env = eq_clause.env in
- match find_positions env sigma t1 t2 with
+ match find_positions env sigma ~keep_proofs ~no_discr:true t1 t2 with
| Inl _ ->
- tclZEROMSG (strbrk"This equality is discriminable. You should use the discriminate tactic to solve the goal.")
+ assert false
| Inr [] ->
let suggestion =
- if !keep_proof_equalities_for_injection then
+ if keep_proofs then
"" else
" You can try to use option Set Keep Proof Equalities." in
tclZEROMSG (strbrk("No information can be deduced from this equality and the injectivity of constructors. This may be because the terms are convertible, or due to pattern matching restrictions in the sort Prop." ^ suggestion))
- | Inr [([],_,_)] when Flags.version_strictly_greater Flags.V8_3 ->
+ | Inr [([],_,_)] ->
tclZEROMSG (str"Nothing to inject.")
| Inr posns ->
inject_at_positions env sigma l2r u eq_clause posns
(tac (clenv_value eq_clause))
let get_previous_hyp_position id gl =
+ let env, sigma = Proofview.Goal.(env gl, sigma gl) in
let rec aux dest = function
- | [] -> raise (RefinerError (NoSuchHyp id))
+ | [] -> raise (RefinerError (env, sigma, NoSuchHyp id))
| d :: right ->
let hyp = Context.Named.Declaration.get_id d in
if Id.equal hyp id then dest else aux (MoveAfter hyp) right
in
- aux MoveLast (Proofview.Goal.hyps (Proofview.Goal.assume gl))
+ aux MoveLast (Proofview.Goal.hyps gl)
-let injEq ?(old=false) with_evars clear_flag ipats =
+let injEq flags ?(old=false) with_evars clear_flag ipats =
(* Decide which compatibility mode to use *)
let ipats_style, l2r, dft_clear_flag, bounded_intro = match ipats with
- | None when not old && use_injection_in_context () ->
+ | None when not old && use_injection_in_context flags ->
Some [], true, true, true
| None -> None, false, false, false
- | _ -> let b = use_injection_pattern_l2r_order () in ipats, b, b, b in
+ | _ -> let b = use_injection_pattern_l2r_order flags in ipats, b, b, b in
(* Built the post tactic depending on compatibility mode *)
let post_tac c n =
match ipats_style with
| Some ipats ->
- Proofview.Goal.enter { enter = begin fun gl ->
- let destopt = match kind_of_term c with
+ Proofview.Goal.enter begin fun gl ->
+ 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 =
@@ -1433,28 +1447,28 @@ let injEq ?(old=false) with_evars clear_flag ipats =
then intro_patterns_bound_to with_evars n destopt ipats
else intro_patterns_to with_evars destopt ipats in
tclTHEN clear_tac intro_tac
- end }
+ end
| None -> tclIDTAC in
- injEqThen post_tac l2r
+ injEqThen (keep_proof_equalities flags) post_tac l2r
-let inj ipats with_evars clear_flag = onEquality with_evars (injEq with_evars clear_flag ipats)
+let inj flags ipats with_evars clear_flag = onEquality with_evars (injEq flags with_evars clear_flag ipats)
-let injClause ipats with_evars = function
- | None -> onNegatedEquality with_evars (injEq with_evars None ipats)
- | Some c -> onInductionArg (inj ipats with_evars) c
+let injClause flags ipats with_evars = function
+ | None -> onNegatedEquality with_evars (injEq flags with_evars None ipats)
+ | Some c -> onInductionArg (inj flags ipats with_evars) c
-let simpleInjClause with_evars = function
- | None -> onNegatedEquality with_evars (injEq ~old:true with_evars None None)
- | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (injEq ~old:true with_evars clear_flag None)) c
+let simpleInjClause flags with_evars = function
+ | None -> onNegatedEquality with_evars (injEq flags ~old:true with_evars None None)
+ | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (injEq flags ~old:true with_evars clear_flag None)) c
-let injConcl = injClause None false None
-let injHyp clear_flag id = injClause None false (Some (clear_flag,ElimOnIdent (Loc.ghost,id)))
+let injConcl flags = injClause flags None false None
+let injHyp flags clear_flag id = injClause flags None false (Some (clear_flag,ElimOnIdent CAst.(make id)))
-let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause =
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+let decompEqThen keep_proofs ntac (lbeq,_,(t,t1,t2) as u) clause =
+ Proofview.Goal.enter begin fun gl ->
let sigma = clause.evd in
let env = Proofview.Goal.env gl in
- match find_positions env sigma t1 t2 with
+ match find_positions env sigma ~keep_proofs ~no_discr:false t1 t2 with
| Inl (cpath, (_,dirn), _) ->
discr_positions env sigma u clause cpath dirn
| Inr [] -> (* Change: do not fail, simplify clear this trivial hyp *)
@@ -1462,21 +1476,21 @@ let decompEqThen ntac (lbeq,_,(t,t1,t2) as u) clause =
| Inr posns ->
inject_at_positions env sigma true u clause posns
(ntac (clenv_value clause))
- end }
+ end
-let dEqThen with_evars ntac = function
- | None -> onNegatedEquality with_evars (decompEqThen (ntac None))
- | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (decompEqThen (ntac clear_flag))) c
+let dEqThen ~keep_proofs with_evars ntac = function
+ | None -> onNegatedEquality with_evars (decompEqThen (use_keep_proofs keep_proofs) (ntac None))
+ | Some c -> onInductionArg (fun clear_flag -> onEquality with_evars (decompEqThen (use_keep_proofs keep_proofs) (ntac clear_flag))) c
-let dEq with_evars =
- dEqThen with_evars (fun clear_flag c x ->
+let dEq ~keep_proofs with_evars =
+ dEqThen ~keep_proofs 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 =
- Proofview.Goal.enter { enter = begin fun gl ->
- let cl = pf_apply make_clenv_binding gl cl NoBindings in
- decompEqThen (fun _ -> tac) data cl
- end }
+let intro_decomp_eq tac data (c, t) =
+ Proofview.Goal.enter begin fun gl ->
+ let cl = pf_apply make_clenv_binding gl (c, t) NoBindings in
+ decompEqThen !keep_proof_equalities_for_injection (fun _ -> tac) data cl
+ end
let _ = declare_intro_decomp_eq intro_decomp_eq
@@ -1512,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 env 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 ->
[]
@@ -1527,11 +1541,10 @@ let decomp_tuple_term env c t =
in decomprec (mkRel 1) c t
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
@@ -1543,16 +1556,16 @@ 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
+ let expected_goal = nf_betaiota env sigma expected_goal in
(* Retype to get universes right *)
let sigma, expected_goal_ty = Typing.type_of env sigma expected_goal in
let sigma, _ = Typing.type_of env sigma body in
- Sigma.Unsafe.of_pair ((body, expected_goal), sigma)
+ (sigma, (body, expected_goal))
(* Like "replace" but decompose dependent equalities *)
(* i.e. if equality is "exists t v = exists u w", and goal is "phi(t,u)", *)
@@ -1560,42 +1573,38 @@ 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.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
let typ = pf_concl gl in
let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in
- let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in
- let tac =
- tclTHENFIRST
+ let (sigma, (typ, expected)) = subst_tuple_term env sigma e1 e2 typ in
+ tclTHEN (Proofview.Unsafe.tclEVARS sigma)
+ (tclTHENFIRST
(tclTHENLIST [
(change_concl typ); (* Put in pattern form *)
(replace_core onConcl l2r eqn)
])
- (change_concl expected) (* Put in normalized form *)
- in
- Sigma (tac, sigma, p)
- end }
+ (change_concl expected)) (* Put in normalized form *)
+ end
let cutSubstInHyp l2r eqn id =
- Proofview.Goal.nf_s_enter { s_enter = begin fun gl ->
+ Proofview.Goal.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
let typ = pf_get_hyp_typ id gl in
let (e1,e2) = if l2r then (e1,e2) else (e2,e1) in
- let Sigma ((typ, expected), sigma, p) = subst_tuple_term env sigma e1 e2 typ in
- let tac =
- tclTHENFIRST
+ let (sigma, (typ, expected)) = subst_tuple_term env sigma e1 e2 typ in
+ tclTHEN (Proofview.Unsafe.tclEVARS sigma)
+ (tclTHENFIRST
(tclTHENLIST [
(change_in_hyp None (make_change_arg typ) (id,InHypTypeOnly));
(replace_core (onHyp id) l2r eqn)
])
- (change_in_hyp None (make_change_arg expected) (id,InHypTypeOnly))
- in
- Sigma (tac, sigma, p)
- end }
+ (change_in_hyp None (make_change_arg expected) (id,InHypTypeOnly)))
+ end
let try_rewrite tac =
Proofview.tclORELSE tac begin function (e, info) -> match e with
@@ -1617,11 +1626,11 @@ let cutRewriteInHyp l2r eqn id = cutRewriteClause l2r eqn (Some id)
let cutRewriteInConcl l2r eqn = cutRewriteClause l2r eqn None
let substClause l2r c cls =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let eq = pf_apply get_type_of gl c in
tclTHENS (cutSubstClause l2r eq cls)
[Proofview.tclUNIT (); exact_no_check c]
- end }
+ end
let rewriteClause l2r c cls = try_rewrite (substClause l2r c cls)
let rewriteInHyp l2r c id = rewriteClause l2r c (Some id)
@@ -1650,8 +1659,7 @@ let regular_subst_tactic = ref true
let _ =
declare_bool_option
- { optsync = true;
- optdepr = false;
+ { optdepr = false;
optname = "more regular behavior of tactic subst";
optkey = ["Regular";"Subst";"Tactic"];
optread = (fun () -> !regular_subst_tactic);
@@ -1666,16 +1674,16 @@ exception FoundHyp of (Id.t * constr * bool)
(* tests whether hyp [c] is [x = t] or [t = x], [x] not occurring in [t] *)
let is_eq_x gl x d =
- let id = get_id d in
+ 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 (get_type d) 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 ->
()
@@ -1683,16 +1691,17 @@ let is_eq_x gl x d =
erase hyp and x; proceed by generalizing all dep hyps *)
let subst_one dep_proof_ok x (hyp,rhs,dir) =
- Proofview.Goal.enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
- let hyps = Proofview.Goal.hyps (Proofview.Goal.assume gl) in
- let concl = Proofview.Goal.concl (Proofview.Goal.assume gl) in
+ let sigma = Tacmach.New.project gl in
+ let hyps = Proofview.Goal.hyps gl in
+ let concl = Proofview.Goal.concl gl in
(* The set of hypotheses using x *)
let dephyps =
List.rev (pi3 (List.fold_right (fun dcl (dest,deps,allhyps) ->
- let id = get_id dcl in
+ 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)
@@ -1701,7 +1710,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
@@ -1711,17 +1720,16 @@ let subst_one dep_proof_ok x (hyp,rhs,dir) =
else
[Proofview.tclUNIT ()]) @
[tclTRY (clear [x; hyp])])
- end }
+ end
(* Look for an hypothesis hyp of the form "x=rhs" or "rhs=x", rewrite
it everywhere, and erase hyp and x; proceed by generalizing all dep hyps *)
let subst_one_var dep_proof_ok x =
- Proofview.Goal.enter { enter = begin fun gl ->
- let gl = Proofview.Goal.assume gl in
- let xval = pf_get_hyp x gl |> get_value in
+ Proofview.Goal.enter begin fun gl ->
+ let decl = pf_get_hyp x gl in
(* If x has a body, simply replace x with body and clear x *)
- if not (Option.is_empty xval) then tclTHEN (unfold_body x) (clear [x]) else
+ if is_local_def decl then tclTHEN (unfold_body x) (clear [x]) else
(* Find a non-recursive definition for x *)
let res =
try
@@ -1729,12 +1737,12 @@ let subst_one_var dep_proof_ok x =
let hyps = Proofview.Goal.hyps gl in
let test hyp _ = is_eq_x gl x hyp in
Context.Named.fold_outside test ~init:() hyps;
- errorlabstrm "Subst"
- (str "Cannot find any non-recursive equality over " ++ pr_id x ++
+ user_err ~hdr:"Subst"
+ (str "Cannot find any non-recursive equality over " ++ Id.print x ++
str".")
with FoundHyp res -> res in
subst_one dep_proof_ok x res
- end }
+ end
let subst_gen dep_proof_ok ids =
tclMAP (subst_one_var dep_proof_ok) ids
@@ -1750,31 +1758,38 @@ type subst_tactic_flags = {
rewrite_dependent_proof : bool
}
-let default_subst_tactic_flags () =
- if Flags.version_strictly_greater Flags.V8_2 then
- { only_leibniz = false; rewrite_dependent_proof = true }
- else
- { only_leibniz = true; rewrite_dependent_proof = false }
+let default_subst_tactic_flags =
+ { only_leibniz = false; rewrite_dependent_proof = true }
+
+let warn_deprecated_simple_subst =
+ CWarnings.create ~name:"deprecated-simple-subst" ~category:"deprecated"
+ (fun () -> strbrk"\"simple subst\" is deprecated")
-let subst_all ?(flags=default_subst_tactic_flags ()) () =
+let subst_all ?(flags=default_subst_tactic_flags) () =
+
+ let () =
+ if flags.only_leibniz || not flags.rewrite_dependent_proof then
+ warn_deprecated_simple_subst ()
+ in
if !regular_subst_tactic then
(* First step: find hypotheses to treat in linear time *)
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 (get_type decl) in
+ 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 (get_id decl)
+ Some (NamedDecl.get_id decl)
| _, Var z when not (is_evaluable env (EvalVarRef z)) ->
- Some (get_id decl)
+ Some (NamedDecl.get_id decl)
| _ ->
None
with Constr_matching.PatternMatchingFailure -> None
@@ -1785,51 +1800,53 @@ let subst_all ?(flags=default_subst_tactic_flags ()) () =
(* Second step: treat equations *)
let process hyp =
- Proofview.Goal.enter { enter = begin fun gl ->
- let gl = Proofview.Goal.assume gl in
+ Proofview.Goal.enter begin fun gl ->
+ 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 |> get_type 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 (dependent 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 (dependent sigma y x) && not (is_evaluable env (EvalVarRef y')) ->
subst_one flags.rewrite_dependent_proof y' (hyp,x,false)
| _ ->
Proofview.tclUNIT ()
- end }
+ end
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
+ Proofview.Goal.enter begin fun gl ->
let ids = find_equations gl in
tclMAP process ids
- end }
+ end
else
(* 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 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
let ids = List.map_filter test hyps in
let ids = List.uniquize ids in
subst_gen flags.rewrite_dependent_proof ids
- end }
+ end
(* Rewrite the first assumption for which a condition holds
and gives the direction of the rewrite *)
@@ -1856,21 +1873,20 @@ let cond_eq_term c t gl =
let rewrite_assumption_cond cond_eq_term cl =
let rec arec hyps gl = match hyps with
- | [] -> error "No such assumption."
+ | [] -> user_err Pp.(str "No such assumption.")
| hyp ::rest ->
- let id = get_id hyp in
+ let id = NamedDecl.get_id hyp in
begin
try
- let dir = cond_eq_term (get_type hyp) gl in
+ let dir = cond_eq_term (NamedDecl.get_type hyp) gl in
general_rewrite_clause dir false (mkVar id,NoBindings) cl
with | Failure _ | UserError _ -> arec rest gl
end
in
- Proofview.Goal.nf_enter { enter = begin fun gl ->
- let gl = Proofview.Goal.lift gl Sigma.Unsafe.le in
+ Proofview.Goal.enter begin fun gl ->
let hyps = Proofview.Goal.hyps gl in
arec hyps gl
- end }
+ end
(* Generalize "subst x" to substitution of subterm appearing as an
equation in the context, but not clearing the hypothesis *)