aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--plugins/ssr/ssrcommon.ml4
-rw-r--r--plugins/ssr/ssrelim.ml2
-rw-r--r--plugins/ssr/ssrequality.ml2
-rw-r--r--plugins/ssr/ssripats.ml2
-rw-r--r--plugins/ssr/ssrtacticals.ml2
-rw-r--r--tactics/tactics.ml10
-rw-r--r--tactics/tactics.mli2
7 files changed, 12 insertions, 12 deletions
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 7cdf05117..d25ecee06 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -745,7 +745,7 @@ let discharge_hyp (id', (id, mode)) gl =
let cl' = Vars.subst_var id (pf_concl gl) in
match pf_get_hyp gl id, mode with
| NamedDecl.LocalAssum (_, t), _ | NamedDecl.LocalDef (_, _, t), "(" ->
- Proofview.V82.of_tactic (Tactics.apply_type (EConstr.of_constr (mkProd (Name id', t, cl')))
+ Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:false (EConstr.of_constr (mkProd (Name id', t, cl')))
[EConstr.of_constr (mkVar id)]) gl
| NamedDecl.LocalDef (_, v, t), _ ->
Proofview.V82.of_tactic
@@ -1179,7 +1179,7 @@ let pf_interp_gen_aux ist gl to_ind ((oclr, occ), t) =
false, pat, EConstr.mkProd (constr_name (project gl) c, pty, Tacmach.pf_concl gl), p, clr,ucst,gl
else CErrors.user_err ?loc:(loc_of_cpattern t) (str "generalized term didn't match")
-let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:false x xs)
let genclrtac cl cs clr =
let tclmyORELSE tac1 tac2 gl =
diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml
index 4e0b44a44..5782a7621 100644
--- a/plugins/ssr/ssrelim.ml
+++ b/plugins/ssr/ssrelim.ml
@@ -28,7 +28,7 @@ module RelDecl = Context.Rel.Declaration
(** The "case" and "elim" tactic *)
-let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:false x xs)
(* TASSI: given the type of an elimination principle, it finds the higher order
* argument (index), it computes it's arity and the arity of the eliminator and
diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml
index 6032ed2af..11ebe4337 100644
--- a/plugins/ssr/ssrequality.ml
+++ b/plugins/ssr/ssrequality.ml
@@ -382,7 +382,7 @@ let is_construct_ref sigma c r =
EConstr.isConstruct sigma c && eq_gr (ConstructRef (fst(EConstr.destConstruct sigma c))) r
let is_ind_ref sigma c r = EConstr.isInd sigma c && eq_gr (IndRef (fst(EConstr.destInd sigma c))) r
-let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:false x xs)
let rwcltac cl rdx dir sr gl =
let n, r_n,_, ucst = pf_abs_evars gl sr in
diff --git a/plugins/ssr/ssripats.ml b/plugins/ssr/ssripats.ml
index 6c325cce4..b3be31b7b 100644
--- a/plugins/ssr/ssripats.ml
+++ b/plugins/ssr/ssripats.ml
@@ -41,7 +41,7 @@ module RelDecl = Context.Rel.Declaration
(* They require guessing the view hints and the number of *)
(* implicits, respectively, which we do by brute force. *)
-let apply_type x xs = Proofview.V82.of_tactic (apply_type x xs)
+let apply_type x xs = Proofview.V82.of_tactic (apply_type ~typecheck:false x xs)
let new_tac = Proofview.V82.of_tactic
diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml
index 5e43c8374..6514b186e 100644
--- a/plugins/ssr/ssrtacticals.ml
+++ b/plugins/ssr/ssrtacticals.ml
@@ -122,7 +122,7 @@ let endclausestac id_map clseq gl_id cl0 gl =
if List.for_all not_hyp' all_ids && not c_hidden then mktac [] gl else
CErrors.user_err (Pp.str "tampering with discharged assumptions of \"in\" tactical")
-let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type x xs)
+let apply_type x xs = Proofview.V82.of_tactic (Tactics.apply_type ~typecheck:false x xs)
let tclCLAUSES ist tac (gens, clseq) gl =
if clseq = InGoal || clseq = InSeqGoal then tac gl else
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 9fded04db..d5160fc9a 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -2157,11 +2157,11 @@ let keep hyps =
and [a1..an:A1..An(a1..an-1)] such that the goal is [G(a1..an)],
this generalizes [hyps |- goal] into [hyps |- T] *)
-let apply_type newcl args =
+let apply_type ~typecheck newcl args =
Proofview.Goal.enter begin fun gl ->
let env = Proofview.Goal.env gl in
let store = Proofview.Goal.extra gl in
- Refine.refine ~typecheck:false begin fun sigma ->
+ Refine.refine ~typecheck begin fun sigma ->
let newcl = nf_betaiota env sigma newcl (* As in former Logic.refine *) in
let (sigma, ev) =
Evarutil.new_evar env sigma ~principal:true ~store newcl in
@@ -2893,7 +2893,7 @@ let generalize_dep ?(with_let=false) c =
let args = Context.Named.to_instance mkVar to_quantify_rev in
tclTHENLIST
[ Proofview.Unsafe.tclEVARS evd;
- apply_type cl'' (if Option.is_empty body then c::args else args);
+ apply_type ~typecheck:false cl'' (if Option.is_empty body then c::args else args);
clear (List.rev tothin')]
end
@@ -2907,7 +2907,7 @@ let generalize_gen_let lconstr = Proofview.Goal.enter begin fun gl ->
let (evd, _) = Typing.type_of env evd newcl in
let map ((_, c, b),_) = if Option.is_empty b then Some c else None in
Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evd)
- (apply_type newcl (List.map_filter map lconstr))
+ (apply_type ~typecheck:false newcl (List.map_filter map lconstr))
end
let new_generalize_gen_let lconstr =
@@ -4282,7 +4282,7 @@ let apply_induction_in_context with_evars hyp0 inhyps elim indvars names induct_
(if isrec then Tacticals.New.tclTHENFIRSTn else Tacticals.New.tclTHENLASTn)
(Tacticals.New.tclTHENLIST [
(* Generalize dependent hyps (but not args) *)
- if deps = [] then Proofview.tclUNIT () else apply_type tmpcl deps_cstr;
+ if deps = [] then Proofview.tclUNIT () else apply_type ~typecheck:false tmpcl deps_cstr;
(* side-conditions in elim (resp case) schemes come last (resp first) *)
induct_tac elim;
Tacticals.New.tclMAP expand_hyp toclear;
diff --git a/tactics/tactics.mli b/tactics/tactics.mli
index 83fc655f1..ca8e0f20c 100644
--- a/tactics/tactics.mli
+++ b/tactics/tactics.mli
@@ -185,7 +185,7 @@ val revert : Id.t list -> unit Proofview.tactic
(** {6 Resolution tactics. } *)
-val apply_type : constr -> constr list -> unit Proofview.tactic
+val apply_type : typecheck:bool -> constr -> constr list -> unit Proofview.tactic
val bring_hyps : named_context -> unit Proofview.tactic
val apply : constr -> unit Proofview.tactic