aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--test-suite/output/clear.out5
-rw-r--r--test-suite/output/clear.v8
-rw-r--r--theories/Init/Tactics.v1
3 files changed, 14 insertions, 0 deletions
diff --git a/test-suite/output/clear.out b/test-suite/output/clear.out
new file mode 100644
index 0000000000..42e3abf26f
--- /dev/null
+++ b/test-suite/output/clear.out
@@ -0,0 +1,5 @@
+1 subgoal
+
+ z := 0 : nat
+ ============================
+ True
diff --git a/test-suite/output/clear.v b/test-suite/output/clear.v
new file mode 100644
index 0000000000..d584cf752e
--- /dev/null
+++ b/test-suite/output/clear.v
@@ -0,0 +1,8 @@
+Module Wish11692.
+
+(* Support for let-in in clear dependent *)
+
+Goal forall x : Prop, let z := 0 in let z' : (fun _ => True) x := I in let y := x in y -> True.
+Proof. intros x z z' y H. clear dependent x. Show. exact I. Qed.
+
+End Wish11692.
diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v
index 6aa1198c08..a4347bbe62 100644
--- a/theories/Init/Tactics.v
+++ b/theories/Init/Tactics.v
@@ -233,6 +233,7 @@ Tactic Notation "clear" "dependent" hyp(h) :=
clear h ||
match goal with
| H : context [ h ] |- _ => depclear H; depclear h
+ | H := context [ h ] |- _ => depclear H; depclear h
end ||
fail "hypothesis to clear is used in the conclusion (maybe indirectly)"
in depclear h.