aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorPierre-Marie Pédrot2020-09-08 10:37:24 +0200
committerPierre-Marie Pédrot2020-09-08 10:37:24 +0200
commitdde607ce50ddcf5f965d4ce222ca50f4d169a2f3 (patch)
tree7f1eb9b6d5af851e776b0a6eb7594eb8b8910c00
parentedb2cefe61d5f5b3230ab33842e61423ca02fbdb (diff)
parente5890700b4574ee9207b2daf6ec95707b9290275 (diff)
Merge PR #12954: Fixes a freshness issue with destruct/induction (see comment in #12944).
Ack-by: RalfJung Ack-by: jashug Reviewed-by: ppedrot
-rw-r--r--tactics/tactics.ml7
-rw-r--r--test-suite/bugs/closed/bug_12944.v12
-rw-r--r--test-suite/success/induct.v10
3 files changed, 24 insertions, 5 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index 5f7e35d205..d33f3a5062 100644
--- a/tactics/tactics.ml
+++ b/tactics/tactics.ml
@@ -3248,13 +3248,10 @@ let rec consume_pattern avoid na isdep gl = let open CAst in function
| {loc;v=IntroForthcoming true}::names when not isdep ->
consume_pattern avoid na isdep gl names
| {loc;v=IntroForthcoming _}::names as fullpat ->
- let avoid = Id.Set.union avoid (explicit_intro_names names) in
(CAst.make ?loc @@ intropattern_of_name gl avoid na, fullpat)
| {loc;v=IntroNaming IntroAnonymous}::names ->
- let avoid = Id.Set.union avoid (explicit_intro_names names) in
(CAst.make ?loc @@ intropattern_of_name gl avoid na, names)
| {loc;v=IntroNaming (IntroFresh id')}::names ->
- let avoid = Id.Set.union avoid (explicit_intro_names names) in
(CAst.make ?loc @@ IntroNaming (IntroIdentifier (new_fresh_id avoid id' gl)), names)
| pat::names -> (pat,names)
@@ -3312,7 +3309,7 @@ let get_recarg_dest (recargdests,tophyp) =
*)
let induct_discharge with_evars dests avoid' tac (avoid,ra) names =
- let avoid = Id.Set.union avoid avoid' in
+ let avoid = Id.Set.union avoid' (Id.Set.union avoid (explicit_intro_names names)) in
let rec peel_tac ra dests names thin =
match ra with
| (RecArg,_,deprec,recvarname) ::
@@ -3320,7 +3317,7 @@ let induct_discharge with_evars dests avoid' tac (avoid,ra) names =
Proofview.Goal.enter begin fun gl ->
let (recpat,names) = match names with
| [{CAst.loc;v=IntroNaming (IntroIdentifier id)} as pat] ->
- let id' = next_ident_away (add_prefix "IH" id) avoid in
+ let id' = new_fresh_id avoid (add_prefix "IH" id) gl in
(pat, [CAst.make @@ IntroNaming (IntroIdentifier id')])
| _ -> consume_pattern avoid (Name recvarname) deprec gl names in
let dest = get_recarg_dest dests in
diff --git a/test-suite/bugs/closed/bug_12944.v b/test-suite/bugs/closed/bug_12944.v
new file mode 100644
index 0000000000..d6720d9906
--- /dev/null
+++ b/test-suite/bugs/closed/bug_12944.v
@@ -0,0 +1,12 @@
+
+Inductive vector A : nat -> Type :=
+ |nil : vector A 0
+ |cons : forall (h:A) (n:nat), vector A n -> vector A (S n).
+
+Global Set Mangle Names.
+
+Lemma vlookup_middle {A n} (v : vector A n) : True.
+Proof.
+ induction v as [|?? IHv].
+ all:exact I.
+Qed.
diff --git a/test-suite/success/induct.v b/test-suite/success/induct.v
index 73fe53c757..a39b17e1f1 100644
--- a/test-suite/success/induct.v
+++ b/test-suite/success/induct.v
@@ -196,3 +196,13 @@ Goal forall m n:nat, n=m.
double induction m n.
Abort.
+(* Mentioned as part of bug #12944 *)
+
+Inductive test : Set := cons : forall (IHv : nat) (v : test), test.
+
+Goal test -> test.
+induction 1 as [? IHv].
+Undo.
+destruct 1 as [? IHv].
+exact IHv. (* Check that the name is granted *)
+Qed.