aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--plugins/micromega/Tauto.v4
-rw-r--r--tactics/tactics.ml6
2 files changed, 8 insertions, 2 deletions
diff --git a/plugins/micromega/Tauto.v b/plugins/micromega/Tauto.v
index 31f55ae9c..458844e1b 100644
--- a/plugins/micromega/Tauto.v
+++ b/plugins/micromega/Tauto.v
@@ -211,7 +211,7 @@ Set Implicit Arguments.
(* BC *)
simpl.
case_eq (deduce t t) ; auto.
- intros until 0.
+ intros *.
case_eq (unsat t0) ; auto.
unfold eval_clause.
rewrite make_conj_cons.
@@ -263,7 +263,7 @@ Set Implicit Arguments.
Proof.
induction cl.
simpl. tauto.
- intros until 0.
+ intros *.
simpl.
assert (HH := add_term_correct env a cl').
case_eq (add_term a cl').
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index df3cca144..834d73bdd 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -1102,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 ->