From 9ebf44d84754adc5b64fcf612c6816c02c80462d Mon Sep 17 00:00:00 2001 From: Benjamin Barenblat Date: Sat, 2 Feb 2019 19:29:23 -0500 Subject: Imported Upstream version 8.9.0 --- tactics/tactics.ml | 296 +++++++++++++++++++++++++++++++---------------------- 1 file changed, 175 insertions(+), 121 deletions(-) (limited to 'tactics/tactics.ml') diff --git a/tactics/tactics.ml b/tactics/tactics.ml index dfdd3237..f060e494 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -43,7 +43,7 @@ open Pretype_errors open Unification open Locus open Locusops -open Misctypes +open Tactypes open Proofview.Notations open Context.Named.Declaration @@ -128,14 +128,14 @@ let unsafe_intro env store decl b = (sigma, mkNamedLambda_or_LetIn decl ev) end -let introduction ?(check=true) id = +let introduction id = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in let hyps = named_context_val (Proofview.Goal.env gl) in let store = Proofview.Goal.extra gl in let env = Proofview.Goal.env gl in - let () = if check && mem_named_context_val id hyps then + let () = if mem_named_context_val id hyps then user_err ~hdr:"Tactics.introduction" (str "Variable " ++ Id.print id ++ str " is already declared.") in @@ -158,9 +158,9 @@ let convert_concl ?(check=true) ty k = let sigma = if check then begin ignore (Typing.unsafe_type_of env sigma ty); - let sigma,b = Reductionops.infer_conv env sigma ty conclty in - if not b then error "Not convertible."; - sigma + match Reductionops.infer_conv env sigma ty conclty with + | None -> error "Not convertible." + | Some sigma -> sigma end else sigma in let (sigma, x) = Evarutil.new_evar env sigma ~principal:true ~store ty in let ans = if k == DEFAULTcast then x else mkCast(x,k,conclty) in @@ -186,11 +186,10 @@ let convert_hyp_no_check = convert_hyp ~check:false let convert_gen pb x y = Proofview.Goal.enter begin fun gl -> - try - let sigma, b = Tacmach.New.pf_apply (Reductionops.infer_conv ~pb) gl x y in - if b then Proofview.Unsafe.tclEVARS sigma - else Tacticals.New.tclFAIL 0 (str "Not convertible") - with (* Reduction.NotConvertible *) _ -> + match Tacmach.New.pf_apply (Reductionops.infer_conv ~pb) gl x y with + | Some sigma -> Proofview.Unsafe.tclEVARS sigma + | None -> Tacticals.New.tclFAIL 0 (str "Not convertible") + | exception _ -> (** FIXME: Sometimes an anomaly is raised from conversion *) Tacticals.New.tclFAIL 0 (str "Not convertible") end @@ -198,32 +197,46 @@ end let convert x y = convert_gen Reduction.CONV x y let convert_leq x y = convert_gen Reduction.CUMUL x y -let clear_dependency_msg env sigma id = function +let clear_in_global_msg = function + | None -> mt () + | Some ref -> str " implicitly in " ++ Printer.pr_global ref + +let clear_dependency_msg env sigma id err inglobal = + let pp = clear_in_global_msg inglobal in + match err with | Evarutil.OccurHypInSimpleClause None -> - Id.print id ++ str " is used in conclusion." + Id.print id ++ str " is used" ++ pp ++ str " in conclusion." | Evarutil.OccurHypInSimpleClause (Some id') -> - Id.print id ++ strbrk " is used in hypothesis " ++ Id.print id' ++ str"." + Id.print id ++ strbrk " is used" ++ pp ++ str " in hypothesis " ++ Id.print id' ++ str"." | Evarutil.EvarTypingBreak ev -> str "Cannot remove " ++ Id.print id ++ strbrk " without breaking the typing of " ++ Printer.pr_existential env sigma ev ++ str"." + | Evarutil.NoCandidatesLeft ev -> + str "Cannot remove " ++ Id.print id ++ str " as it would leave the existential " ++ + Printer.pr_existential_key sigma ev ++ str" without candidates." -let error_clear_dependency env sigma id err = - user_err (clear_dependency_msg env sigma id err) +let error_clear_dependency env sigma id err inglobal = + user_err (clear_dependency_msg env sigma id err inglobal) -let replacing_dependency_msg env sigma id = function +let replacing_dependency_msg env sigma id err inglobal = + let pp = clear_in_global_msg inglobal in + match err with | Evarutil.OccurHypInSimpleClause None -> - str "Cannot change " ++ Id.print id ++ str ", it is used in conclusion." + str "Cannot change " ++ Id.print id ++ str ", it is used" ++ pp ++ str " in conclusion." | Evarutil.OccurHypInSimpleClause (Some id') -> str "Cannot change " ++ Id.print id ++ - strbrk ", it is used in hypothesis " ++ Id.print id' ++ str"." + strbrk ", it is used" ++ pp ++ str " in hypothesis " ++ Id.print id' ++ str"." | Evarutil.EvarTypingBreak ev -> str "Cannot change " ++ Id.print id ++ strbrk " without breaking the typing of " ++ Printer.pr_existential env sigma ev ++ str"." + | Evarutil.NoCandidatesLeft ev -> + str "Cannot change " ++ Id.print id ++ str " as it would leave the existential " ++ + Printer.pr_existential_key sigma ev ++ str" without candidates." -let error_replacing_dependency env sigma id err = - user_err (replacing_dependency_msg env sigma id err) +let error_replacing_dependency env sigma id err inglobal = + user_err (replacing_dependency_msg env sigma id err inglobal) (* This tactic enables the user to remove hypotheses from the signature. * Some care is taken to prevent him from removing variables that are @@ -239,13 +252,12 @@ let clear_gen fail = function let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let concl = Proofview.Goal.concl gl in - let evdref = ref sigma in - let (hyps, concl) = - try clear_hyps_in_evi env evdref (named_context_val env) concl ids - with Evarutil.ClearDependencyError (id,err) -> fail env sigma id err + let (sigma, hyps, concl) = + try clear_hyps_in_evi env sigma (named_context_val env) concl ids + with Evarutil.ClearDependencyError (id,err,inglobal) -> fail env sigma id err inglobal in let env = reset_with_named_context hyps env in - Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) + Proofview.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Refine.refine ~typecheck:false begin fun sigma -> Evarutil.new_evar env sigma ~principal:true concl end) @@ -423,11 +435,10 @@ let get_previous_hyp_position env sigma id = let clear_hyps2 env sigma ids sign t cl = try - let evdref = ref (Evd.clear_metas sigma) in - let (hyps,t,cl) = Evarutil.clear_hyps2_in_evi env evdref sign t cl ids in - (hyps, t, cl, !evdref) - with Evarutil.ClearDependencyError (id,err) -> - error_replacing_dependency env sigma id err + let sigma = Evd.clear_metas sigma in + Evarutil.clear_hyps2_in_evi env sigma sign t cl ids + with Evarutil.ClearDependencyError (id,err,inglobal) -> + error_replacing_dependency env sigma id err inglobal let internal_cut_gen ?(check=true) dir replace id t = Proofview.Goal.enter begin fun gl -> @@ -439,7 +450,7 @@ let internal_cut_gen ?(check=true) dir replace id t = let sign',t,concl,sigma = if replace then let nexthyp = get_next_hyp_position env sigma id (named_context_of_val sign) in - let sign',t,concl,sigma = clear_hyps2 env sigma (Id.Set.singleton id) sign t concl in + let sigma,sign',t,concl = clear_hyps2 env sigma (Id.Set.singleton id) sign t concl in let sign' = insert_decl_in_named_context sigma (LocalAssum (id,t)) nexthyp sign' in sign',t,concl,sigma else @@ -557,20 +568,7 @@ let mutual_fix f n rest j = Proofview.Goal.enter begin fun gl -> end end -let warning_nameless_fix = - CWarnings.create ~name:"nameless-fix" ~category:"deprecated" Pp.(fun () -> - str "fix/cofix without a name are deprecated, please use the named version.") - -let fix ido n = match ido with - | None -> - warning_nameless_fix (); - Proofview.Goal.enter begin fun gl -> - let name = Proof_global.get_current_proof_name () in - let id = new_fresh_id Id.Set.empty name gl in - mutual_fix id n [] 0 - end - | Some id -> - mutual_fix id n [] 0 +let fix id n = mutual_fix id n [] 0 let rec check_is_mutcoind env sigma cl = let b = whd_all env sigma cl in @@ -613,16 +611,7 @@ let mutual_cofix f others j = Proofview.Goal.enter begin fun gl -> end end -let cofix ido = match ido with - | None -> - warning_nameless_fix (); - Proofview.Goal.enter begin fun gl -> - let name = Proof_global.get_current_proof_name () in - let id = new_fresh_id Id.Set.empty name gl in - mutual_cofix id [] 0 - end - | Some id -> - mutual_cofix id [] 0 +let cofix id = mutual_cofix id [] 0 (**************************************************************) (* Reduction and conversion tactics *) @@ -802,9 +791,9 @@ let e_change_in_hyp redfun (id,where) = (convert_hyp c) end -type change_arg = Ltac_pretype.patvar_map -> evar_map -> evar_map * EConstr.constr +type change_arg = Ltac_pretype.patvar_map -> env -> evar_map -> evar_map * EConstr.constr -let make_change_arg c pats sigma = (sigma, replace_vars (Id.Map.bindings pats) c) +let make_change_arg c pats env sigma = (sigma, replace_vars (Id.Map.bindings pats) c) let check_types env sigma mayneedglobalcheck deep newc origc = let t1 = Retyping.get_type_of env sigma newc in @@ -812,15 +801,15 @@ let check_types env sigma mayneedglobalcheck deep newc origc = let t2 = Retyping.get_type_of env sigma origc in let sigma, t2 = Evarsolve.refresh_universes ~onlyalg:true (Some false) env sigma t2 in - let sigma, b = infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 in - if not b then + match infer_conv ~pb:Reduction.CUMUL env sigma t1 t2 with + | None -> if 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 + | Some sigma -> sigma end else if not (isSort sigma (whd_all env sigma t1)) then @@ -829,11 +818,11 @@ let check_types env sigma mayneedglobalcheck deep newc origc = (* Now we introduce different instances of the previous tacticals *) let change_and_check cv_pb mayneedglobalcheck deep t env sigma c = - let (sigma, t') = t sigma in + let (sigma, t') = t env sigma in let sigma = check_types env sigma mayneedglobalcheck deep t' c in - let sigma, b = infer_conv ~pb:cv_pb env sigma t' c in - if not b then user_err ~hdr:"convert-check-hyp" (str "Not convertible."); - (sigma, t') + match infer_conv ~pb:cv_pb env sigma t' c with + | None -> user_err ~hdr:"convert-check-hyp" (str "Not convertible."); + | Some sigma -> (sigma, t') (* Use cumulativity only if changing the conclusion not a subterm *) let change_on_subterm cv_pb deep t where env sigma c = @@ -971,6 +960,11 @@ let rec intro_then_gen name_flag move_flag force_flag dep_flag tac = | 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 + | Evar ev when force_flag -> + let sigma, t = Evardefine.define_evar_as_product sigma ev in + Tacticals.New.tclTHEN + (Proofview.Unsafe.tclEVARS sigma) + (intro_then_gen name_flag move_flag force_flag dep_flag tac) | _ -> begin if not force_flag then Proofview.tclZERO (RefinerError (env, sigma, IntroNeedsProduct)) (* Note: red_in_concl includes betaiotazeta and this was like *) @@ -1108,7 +1102,13 @@ let msg_quantified_hypothesis = function pr_nth n ++ str " non dependent hypothesis" +let warn_deprecated_intros_until_0 = + CWarnings.create ~name:"deprecated-intros-until-0" ~category:"tactics" + (fun () -> + strbrk"\"intros until 0\" is deprecated, use \"intros *\"; instead of \"induction 0\" and \"destruct 0\" use explicitly a name.\"") + let depth_of_quantified_hypothesis red h gl = + if h = AnonHyp 0 then warn_deprecated_intros_until_0 (); match lookup_hypothesis_as_renamed_gen red h gl with | Some depth -> depth | None -> @@ -1159,6 +1159,18 @@ let tactic_infer_flags with_evar = { Pretyping.fail_evar = not with_evar; Pretyping.expand_evars = true } +type evars_flag = bool (* true = pose evars false = fail on evars *) +type rec_flag = bool (* true = recursive false = not recursive *) +type advanced_flag = bool (* true = advanced false = basic *) +type clear_flag = bool option (* true = clear hyp, false = keep hyp, None = use default *) + +type 'a core_destruction_arg = + | ElimOnConstr of 'a + | ElimOnIdent of lident + | ElimOnAnonHyp of int + +type 'a destruction_arg = + clear_flag * 'a core_destruction_arg let onOpenInductionArg env sigma tac = function | clear_flag,ElimOnConstr f -> @@ -1251,7 +1263,6 @@ 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 " ++ Id.print id ++ str".") @@ -1261,7 +1272,7 @@ let check_unresolved_evars_of_metas sigma clenv = (* Refiner.pose_all_metas_as_evars are resolved *) List.iter (fun (mv,b) -> match b with | Clval (_,(c,_),_) -> - (match Constr.kind c.rebus with + (match Constr.kind (EConstr.Unsafe.to_constr c.rebus) with | Evar (evk,_) when Evd.is_undefined clenv.evd evk && not (Evd.mem sigma evk) -> error_uninstantiated_metas (mkMeta mv) clenv @@ -1281,7 +1292,7 @@ let do_replace id = function let clenv_refine_in ?(sidecond_first=false) with_evars ?(with_classes=true) targetid id sigma0 clenv tac = - let clenv = Clenvtac.clenv_pose_dependent_evars with_evars clenv in + let clenv = Clenvtac.clenv_pose_dependent_evars ~with_evars clenv in let clenv = if with_classes then { clenv with evd = Typeclasses.resolve_typeclasses @@ -1438,9 +1449,7 @@ let is_nonrec mind = (Global.lookup_mind (fst mind)).mind_finite == Declarations 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 + Tacmach.New.pf_apply Evd.fresh_global gl gr let find_eliminator c gl = let ((ind,u),t) = Tacmach.New.pf_reduce_to_quantified_ind gl (Tacmach.New.pf_unsafe_type_of gl c) in @@ -1572,9 +1581,10 @@ let make_projection env sigma params cstr sign elim i n c u = | Some proj -> let args = Context.Rel.to_extended_vect mkRel 0 sign in let proj = - if Environ.is_projection proj env then + match Recordops.find_primitive_projection proj with + | Some proj -> mkProj (Projection.make proj false, mkApp (c, args)) - else + | None -> mkApp (mkConstU (proj,u), Array.append (Array.of_list params) [|mkApp (c, args)|]) in @@ -1639,13 +1649,11 @@ let tclORELSEOPT t k = Proofview.tclZERO ~info e | Some tac -> tac) -let general_apply with_delta with_destruct with_evars clear_flag - {CAst.loc;v=(c,lbind : EConstr.constr with_bindings)} = +let general_apply ?(respect_opaque=false) with_delta with_destruct with_evars + clear_flag {CAst.loc;v=(c,lbind : EConstr.constr with_bindings)} = Proofview.Goal.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. *) @@ -1654,7 +1662,12 @@ let general_apply with_delta with_destruct with_evars clear_flag Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in - + let ts = + if respect_opaque then Conv_oracle.get_transp_state (oracle env) + else full_transparent_state + in + let flags = + if with_delta then default_unify_flags () else default_no_delta_unify_flags ts in let thm_ty0 = nf_betaiota env sigma (Retyping.get_type_of env sigma c) in let try_apply thm_ty nprod = try @@ -1720,14 +1733,14 @@ let rec apply_with_bindings_gen b e = function (general_apply b b e k cb) (apply_with_bindings_gen b e cbl) -let apply_with_delayed_bindings_gen b e l = +let apply_with_delayed_bindings_gen b e l = let one k {CAst.loc;v=f} = Proofview.Goal.enter begin fun gl -> let sigma = Tacmach.New.project gl in let env = Proofview.Goal.env gl in let (sigma, cb) = f env sigma in Tacticals.New.tclWITHHOLES e - (general_apply b b e k CAst.(make ?loc cb)) sigma + (general_apply ~respect_opaque:(not b) b b e k CAst.(make ?loc cb)) sigma end in let rec aux = function @@ -1802,14 +1815,12 @@ let apply_in_once_main flags innerclause env sigma (loc,d,lbind) = in aux (make_clenv_binding env sigma (d,thm) lbind) -let apply_in_once sidecond_first with_delta with_destruct with_evars naming - id (clear_flag,{ CAst.loc; v= d,lbind}) tac = +let apply_in_once ?(respect_opaque = false) sidecond_first with_delta + with_destruct with_evars naming id (clear_flag,{ CAst.loc; v= d,lbind}) tac = let open Context.Rel.Declaration in Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env 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 let t' = Tacmach.New.pf_get_hyp_typ id gl in let innerclause = mk_clenv_from_env env sigma (Some 0) (mkVar id,t') in let targetid = find_name true (LocalAssum (Anonymous,t')) naming gl in @@ -1817,6 +1828,12 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in + let ts = + if respect_opaque then Conv_oracle.get_transp_state (oracle env) + else full_transparent_state + in + let flags = + if with_delta then default_unify_flags () else default_no_delta_unify_flags ts in try let clause = apply_in_once_main flags innerclause env sigma (loc,c,lbind) in clenv_refine_in ~sidecond_first with_evars targetid id sigma clause @@ -1836,14 +1853,14 @@ let apply_in_once sidecond_first with_delta with_destruct with_evars naming aux [] with_destruct d end -let apply_in_delayed_once sidecond_first with_delta with_destruct with_evars naming - id (clear_flag,{CAst.loc;v=f}) tac = +let apply_in_delayed_once ?(respect_opaque = false) sidecond_first with_delta + with_destruct with_evars naming id (clear_flag,{CAst.loc;v=f}) tac = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = Tacmach.New.project gl in let (sigma, c) = f env sigma in Tacticals.New.tclWITHHOLES with_evars - (apply_in_once sidecond_first with_delta with_destruct with_evars + (apply_in_once ~respect_opaque sidecond_first with_delta with_destruct with_evars naming id (clear_flag,CAst.(make ?loc c)) tac) sigma end @@ -1911,8 +1928,8 @@ let cast_no_check cast c = exact_no_check (mkCast (c, cast, concl)) end -let vm_cast_no_check c = cast_no_check Term.VMcast c -let native_cast_no_check c = cast_no_check Term.NATIVEcast c +let vm_cast_no_check c = cast_no_check VMcast c +let native_cast_no_check c = cast_no_check NATIVEcast c let exact_proof c = let open Tacmach.New in @@ -1935,16 +1952,19 @@ let assumption = let t = NamedDecl.get_type decl in let concl = Proofview.Goal.concl gl in let sigma = Tacmach.New.project gl in - let (sigma, is_same_type) = - if only_eq then (sigma, EConstr.eq_constr sigma t concl) + let ans = + if only_eq then + if EConstr.eq_constr sigma t concl then Some sigma + else None else let env = Proofview.Goal.env gl in infer_conv env sigma t concl in - if is_same_type then + match ans with + | Some sigma -> (Proofview.Unsafe.tclEVARS sigma) <*> exact_no_check (mkVar (NamedDecl.get_id decl)) - else arec gl only_eq rest + | None -> arec gl only_eq rest in let assumption_tac gl = let hyps = Proofview.Goal.hyps gl in @@ -1964,24 +1984,22 @@ let on_the_bodies = function exception DependsOnBody of Id.t option let check_is_type env sigma ty = - let evdref = ref sigma in try - let _ = Typing.e_sort_of env evdref ty in - !evdref + let sigma, _ = Typing.sort_of env sigma ty in + sigma with e when CErrors.noncritical e -> raise (DependsOnBody None) let check_decl env sigma decl = let open Context.Named.Declaration in let ty = NamedDecl.get_type decl in - let evdref = ref sigma in try - let _ = Typing.e_sort_of env evdref ty in - let _ = match decl with - | LocalAssum _ -> () - | LocalDef (_,c,_) -> Typing.e_check env evdref c ty + let sigma, _ = Typing.sort_of env sigma ty in + let sigma = match decl with + | LocalAssum _ -> sigma + | LocalDef (_,c,_) -> Typing.check env sigma c ty in - !evdref + sigma with e when CErrors.noncritical e -> let id = NamedDecl.get_id decl in raise (DependsOnBody (Some id)) @@ -2252,7 +2270,7 @@ let intro_or_and_pattern ?loc with_evars bracketed ll thin tac id = let c = mkVar id in let t = Tacmach.New.pf_unsafe_type_of gl c in let (ind,t) = Tacmach.New.pf_reduce_to_quantified_ind gl t in - let branchsigns = compute_constructor_signatures false ind in + let branchsigns = compute_constructor_signatures ~rec_flag:false ind in let nv_with_let = Array.map List.length branchsigns in let ll = fix_empty_or_and_pattern (Array.length branchsigns) ll in let ll = get_and_check_or_and_pattern ?loc ll branchsigns in @@ -2532,11 +2550,11 @@ let assert_as first hd ipat t = (* apply in as *) -let general_apply_in sidecond_first with_delta with_destruct with_evars - id lemmas ipat = +let general_apply_in ?(respect_opaque=false) sidecond_first with_delta + with_destruct with_evars id lemmas ipat = let tac (naming,lemma) tac id = - apply_in_delayed_once sidecond_first with_delta with_destruct with_evars - naming id lemma tac in + apply_in_delayed_once ~respect_opaque sidecond_first with_delta + with_destruct with_evars naming id lemma tac in Proofview.Goal.enter begin fun gl -> let destopt = if with_evars then MoveLast (* evars would depend on the whole context *) @@ -2567,7 +2585,7 @@ let apply_in simple with_evars id lemmas ipat = general_apply_in false simple simple with_evars id lemmas ipat let apply_delayed_in simple with_evars id lemmas ipat = - general_apply_in false simple simple with_evars id lemmas ipat + general_apply_in ~respect_opaque:true false simple simple with_evars id lemmas ipat (*****************************) (* Tactics abstracting terms *) @@ -2605,9 +2623,7 @@ 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) = Evd.fresh_global env sigma eqdata.eq in - let eq = EConstr.of_constr eq in let (sigma, refl) = Evd.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 @@ -2663,9 +2679,7 @@ 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) = Evd.fresh_global env sigma eqdata.eq in - let eq = EConstr.of_constr eq in let (sigma, refl) = Evd.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 @@ -3003,8 +3017,24 @@ let unfold_body x = end end +let warn_cannot_remove_as_expected = + CWarnings.create ~name:"cannot-remove-as-expected" ~category:"tactics" + (fun (id,inglobal) -> + let pp = match inglobal with + | None -> mt () + | Some ref -> str ", it is used implicitly in " ++ Printer.pr_global ref in + str "Cannot remove " ++ Id.print id ++ pp ++ str ".") + +let clear_for_destruct ids = + Proofview.tclORELSE + (clear_gen (fun env sigma id err inglobal -> raise (ClearDependencyError (id,err,inglobal))) ids) + (function + | ClearDependencyError (id,err,inglobal),_ -> warn_cannot_remove_as_expected (id,inglobal); Proofview.tclUNIT () + | e -> iraise e) + (* Either unfold and clear if defined or simply clear if not a definition *) -let expand_hyp id = Tacticals.New.tclTRY (unfold_body id) <*> clear [id] +let expand_hyp id = + Tacticals.New.tclTRY (unfold_body id) <*> clear_for_destruct [id] (*****************************) (* High-level induction *) @@ -3420,7 +3450,7 @@ let cook_sign hyp0_opt inhyps indvars env sigma = type elim_scheme = { elimc: constr with_bindings option; elimt: types; - indref: global_reference option; + indref: GlobRef.t option; params: rel_context; (* (prm1,tprm1);(prm2,tprm2)...(prmp,tprmp) *) nparams: int; (* number of parameters *) predicates: rel_context; (* (Qq, (Tq_1 -> Tq_2 ->...-> Tq_nq)), (Q1,...) *) @@ -3782,7 +3812,10 @@ let specialize_eqs id = let ty = Tacmach.New.pf_get_hyp_typ id gl in let evars = ref (Proofview.Goal.sigma gl) in let unif env evars c1 c2 = - compare_upto_variables !evars c1 c2 && Evarconv.e_conv env evars c1 c2 + compare_upto_variables !evars c1 c2 && + (match Evarconv.conv env !evars c1 c2 with + | Some sigma -> evars := sigma; true + | None -> false) in let rec aux in_eqs ctx acc ty = match EConstr.kind !evars ty with @@ -3807,7 +3840,8 @@ let specialize_eqs id = | _ -> if in_eqs then acc, in_eqs, ctx, ty else - let e = e_new_evar (push_rel_context ctx env) evars t in + let sigma, e = Evarutil.new_evar (push_rel_context ctx env) !evars t in + evars := sigma; aux false (LocalDef (na,e,t) :: ctx) (mkApp (lift 1 acc, [| mkRel 1 |])) b) | t -> acc, in_eqs, ctx, ty in @@ -4176,7 +4210,7 @@ let induction_tac with_evars params indvars elim = let elimclause' = recolle_clenv i params indvars elimclause gl in (* one last resolution (useless?) *) let resolved = clenv_unique_resolver ~flags:(elim_flags ()) elimclause' gl in - Clenvtac.clenv_refine with_evars resolved + Clenvtac.clenv_refine ~with_evars resolved end (* Apply induction "in place" taking into account dependent @@ -4334,7 +4368,7 @@ let check_expected_type env sigma (elimc,bl) elimt = 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 sigma cl.cl_concl in - fun t -> Evarconv.e_cumul env (ref sigma) t u + fun t -> Option.has_some (Evarconv.cumul env sigma t u) let check_enough_applied env sigma elim = (* A heuristic to decide whether the induction arg is enough applied *) @@ -4927,9 +4961,9 @@ let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK = let evd, ctx, concl = (* FIXME: should be done only if the tactic succeeds *) - let evd, nf = nf_evars_and_universes !evdref in + let evd = Evd.minimize_universes !evdref in let ctx = Evd.universe_context_set evd in - evd, ctx, nf concl + evd, ctx, Evarutil.nf_evars_universes evd concl in let concl = EConstr.of_constr concl in let solve_tac = tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac) in @@ -4973,7 +5007,7 @@ let cache_term_by_tactic_then ~opaque ?(goal_type=None) id gk tac tacK = 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 - let effs = add_private eff + let effs = concat_private eff Entries.(snd (Future.force const.const_entry_body)) in let solve = Proofview.tclEFFECTS effs <*> @@ -5007,6 +5041,26 @@ let tclABSTRACT ?(opaque=true) name_op tac = else name_op_to_name name_op (DefinitionBody Definition) "_subterm" in abstract_subproof ~opaque s gk tac +let constr_eq ~strict x y = + let fail = Tacticals.New.tclFAIL 0 (str "Not equal") in + let fail_universes = Tacticals.New.tclFAIL 0 (str "Not equal (due to universes)") in + Proofview.Goal.enter begin fun gl -> + let env = Tacmach.New.pf_env gl in + let evd = Tacmach.New.project gl in + match EConstr.eq_constr_universes env evd x y with + | Some csts -> + let csts = UnivProblem.to_constraints ~force_weak:false (Evd.universes evd) csts in + if strict then + if Evd.check_constraints evd csts then Proofview.tclUNIT () + else fail_universes + else + (match Evd.add_constraints evd csts with + | evd -> Proofview.Unsafe.tclEVARS evd + | exception Univ.UniverseInconsistency _ -> + fail_universes) + | None -> fail + end + let unify ?(state=full_transparent_state) x y = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in -- cgit v1.2.3