From 11eacd8b0ef0ab77fdfe25267732adfdd01f32b2 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Wed, 23 May 2018 16:47:07 +0200 Subject: [PATCH] A stronger version of `cinv_open`. This version allows one to either close or cancel the invariant after opening it. --- .../base_logic/lib/cancelable_invariants.v | 32 ++++++++++++------- 1 file changed, 21 insertions(+), 11 deletions(-) diff --git a/theories/base_logic/lib/cancelable_invariants.v b/theories/base_logic/lib/cancelable_invariants.v index 5e7f386f6..6d18584e5 100644 --- a/theories/base_logic/lib/cancelable_invariants.v +++ b/theories/base_logic/lib/cancelable_invariants.v @@ -70,6 +70,21 @@ Section proofs. iIntros "!>". iExists P. iSplit; last done. iIntros "!# !>"; iSplit; auto. Qed. + Lemma cinv_open_strong E N γ p P : + ↑N ⊆ E → + cinv N γ P -∗ cinv_own γ p ={E,E∖↑N}=∗ + ▷ P ∗ cinv_own γ p ∗ (▷ P ∨ cinv_own γ 1 ={E∖↑N,E}=∗ True). + Proof. + iIntros (?) "#Hinv Hγ". iDestruct "Hinv" as (P') "[#HP' Hinv]". + iInv N as "[HP | >Hγ']" "Hclose". + - iIntros "!> {$Hγ}". iSplitL "HP". + + iNext. iApply "HP'". done. + + iIntros "[HP|Hγ]". + * iApply "Hclose". iLeft. iNext. by iApply "HP'". + * iApply "Hclose". iRight. by iNext. + - iDestruct (cinv_own_1_l with "Hγ' Hγ") as %[]. + Qed. + Lemma cinv_alloc E N P : ▷ P ={E}=∗ ∃ γ, cinv N γ P ∗ cinv_own γ 1. Proof. iIntros "HP". iMod (cinv_alloc_strong ∅ E N) as (γ _) "[Hγ Halloc]". @@ -78,23 +93,18 @@ Section proofs. Lemma cinv_cancel E N γ P : ↑N ⊆ E → cinv N γ P -∗ cinv_own γ 1 ={E}=∗ ▷ P. Proof. - iIntros (?) "#Hinv Hγ". iDestruct "Hinv" as (P') "[#HP' Hinv]". - iInv N as "[HP|>Hγ']" "Hclose". - - iMod ("Hclose" with "[Hγ]") as "_"; first by eauto. iModIntro. iNext. - iApply "HP'". done. - - iDestruct (cinv_own_1_l with "Hγ Hγ'") as %[]. + iIntros (?) "#Hinv Hγ". + iMod (cinv_open_strong with "Hinv Hγ") as "($ & Hγ & H)"; first done. + iApply "H". by iRight. Qed. Lemma cinv_open E N γ p P : ↑N ⊆ E → cinv N γ P -∗ cinv_own γ p ={E,E∖↑N}=∗ ▷ P ∗ cinv_own γ p ∗ (▷ P ={E∖↑N,E}=∗ True). Proof. - iIntros (?) "#Hinv Hγ". iDestruct "Hinv" as (P') "[#HP' Hinv]". - iInv N as "[HP | >Hγ']" "Hclose". - - iIntros "!> {$Hγ}". iSplitL "HP". - + iNext. iApply "HP'". done. - + iIntros "HP". iApply "Hclose". iLeft. iNext. by iApply "HP'". - - iDestruct (cinv_own_1_l with "Hγ' Hγ") as %[]. + iIntros (?) "#Hinv Hγ". + iMod (cinv_open_strong with "Hinv Hγ") as "($ & $ & H)"; first done. + iIntros "!> HP". iApply "H"; auto. Qed. End proofs. -- GitLab