aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHugo Herbelin2020-08-30 21:44:36 +0200
committerHugo Herbelin2020-08-31 12:33:52 +0200
commit4c3064c6e679cfbe4ce13a76b50a17bcfffe4d71 (patch)
tree1188f5ae30cf85463cd54e9164eefb6b0dc8166c
parent94d9fe2b6cb06fe7f862683d8337433e18314001 (diff)
Fixes a freshness issue with induction (see comment in #12944).
The names computed in consume_pattern were lost when calling intros_pattern_core. Moreover, the computation of names to avoid was done several time. We compute it once for all.
-rw-r--r--tactics/tactics.ml5
-rw-r--r--test-suite/success/induct.v10
2 files changed, 11 insertions, 4 deletions
diff --git a/tactics/tactics.ml b/tactics/tactics.ml
index eb7b7e363f..00bd56d8ba 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) ::
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.