Skip to content
Snippets Groups Projects
Commit fda80965 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Fix regression caused by e17ac4ad and another similar bug.

parent c0253c3e
No related branches found
No related tags found
No related merge requests found
...@@ -535,16 +535,17 @@ a goal [P] for non-dependent arguments [x_i : P]. *) ...@@ -535,16 +535,17 @@ a goal [P] for non-dependent arguments [x_i : P]. *)
Tactic Notation "iIntoValid" open_constr(t) := Tactic Notation "iIntoValid" open_constr(t) :=
let rec go t := let rec go t :=
let tT := type of t in let tT := type of t in
first lazymatch eval hnf in tT with
[apply (proj1 (_ : AsValid tT _) t) | ?P ?Q => let H := fresh in assert P as H; [|go uconstr:(t H); clear H]
|lazymatch eval hnf in tT with | _ : ?T, _ =>
| ?P ?Q => let H := fresh in assert P as H; [|go uconstr:(t H); clear H] (* Put [T] inside an [id] to avoid TC inference from being invoked. *)
| _ : ?T, _ => (* This is a workarround for Coq bug #4969. *)
(* Put [T] inside an [id] to avoid TC inference from being invoked. *) let e := fresh in evar (e:id T);
(* This is a workarround for Coq bug #4969. *) let e' := eval unfold e in e in clear e; go (t e')
let e := fresh in evar (e:id T); | _ =>
let e' := eval unfold e in e in clear e; go (t e') let tT' := eval cbv zeta in tT in apply (proj1 (_ : AsValid tT' _) t)
end] in || fail "iPoseProof: not a uPred"
end in
go t. go t.
(* The tactic [tac] is called with a temporary fresh name [H]. The argument (* The tactic [tac] is called with a temporary fresh name [H]. The argument
......
...@@ -201,6 +201,18 @@ Lemma test_True_intros : (True : uPred M) -∗ True. ...@@ -201,6 +201,18 @@ Lemma test_True_intros : (True : uPred M) -∗ True.
Proof. Proof.
iIntros "?". done. iIntros "?". done.
Qed. Qed.
Lemma test_iPoseProof_let P Q :
(let R := True%I in R P Q)
P Q.
Proof.
iIntros (help) "HP".
iPoseProof (help with "[$HP]") as "?". done.
Qed.
Lemma test_iIntros_let P :
Q, let R := True%I in P -∗ R -∗ Q -∗ P Q.
Proof. iIntros (Q R) "$ HR $". Qed.
End tests. End tests.
Section more_tests. Section more_tests.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment