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

Fix regression caused by 0b462620.

This regression was caused by a bug in handling spec patterns.
parent 6fcaabd1
No related branches found
No related tags found
No related merge requests found
...@@ -900,6 +900,16 @@ Proof. ...@@ -900,6 +900,16 @@ Proof.
(into_persistent _ R) sep_elim_r persistently_and_affinely_sep_l wand_elim_r. (into_persistent _ R) sep_elim_r persistently_and_affinely_sep_l wand_elim_r.
Qed. Qed.
(* A special version of [tac_assumption] that does not do any of the
[FromAssumption] magic. *)
Lemma tac_specialize_persistent_helper_done Δ i q P :
envs_lookup i Δ = Some (q,P)
envs_entails Δ (bi_absorbingly P).
Proof.
rewrite envs_entails_eq /bi_absorbingly=> /envs_lookup_sound=> ->.
rewrite affinely_persistently_if_elim comm. f_equiv; auto.
Qed.
Lemma tac_revert Δ Δ' i p P Q : Lemma tac_revert Δ Δ' i p P Q :
envs_lookup_delete true i Δ = Some (p,P,Δ') envs_lookup_delete true i Δ = Some (p,P,Δ')
envs_entails Δ' ((if p then P else P)%I -∗ Q) envs_entails Δ' ((if p then P else P)%I -∗ Q)
......
...@@ -615,7 +615,9 @@ Tactic Notation "iSpecializeCore" open_constr(H) ...@@ -615,7 +615,9 @@ Tactic Notation "iSpecializeCore" open_constr(H)
(* FIXME: do something reasonable when the BI is not affine *) (* FIXME: do something reasonable when the BI is not affine *)
eapply tac_specialize_persistent_helper with _ H _ _ _ _; eapply tac_specialize_persistent_helper with _ H _ _ _ _;
[env_reflexivity || fail "iSpecialize:" H "not found" [env_reflexivity || fail "iSpecialize:" H "not found"
|iSpecializePat H pat; last (iExact H) |iSpecializePat H pat;
[..
|eapply tac_specialize_persistent_helper_done with H _; env_reflexivity]
|apply _ || |apply _ ||
let Q := match goal with |- IntoPersistent _ ?Q _ => Q end in let Q := match goal with |- IntoPersistent _ ?Q _ => Q end in
fail "iSpecialize:" Q "not persistent" fail "iSpecialize:" Q "not persistent"
......
...@@ -94,6 +94,11 @@ Proof. ...@@ -94,6 +94,11 @@ Proof.
done. done.
Qed. Qed.
Lemma test_iAssert_persistently P : P -∗ True.
Proof.
iIntros "HP". iAssert ( P)%I with "[# //]" as "#H". done.
Qed.
Lemma test_iSpecialize_auto_frame P Q R : Lemma test_iSpecialize_auto_frame P Q R :
(P -∗ True -∗ True -∗ Q -∗ R) -∗ P -∗ Q -∗ R. (P -∗ True -∗ True -∗ Q -∗ R) -∗ P -∗ Q -∗ R.
Proof. iIntros "H ? HQ". by iApply ("H" with "[$]"). Qed. Proof. iIntros "H ? HQ". by iApply ("H" with "[$]"). Qed.
......
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