diff --git a/theories/lifetime/borrow.v b/theories/lifetime/borrow.v index 04c1c297783d9388089c9e7fc51669659ede0914..b02e6741ae6d467479b79060bc8b46cfe6065096 100644 --- a/theories/lifetime/borrow.v +++ b/theories/lifetime/borrow.v @@ -13,7 +13,7 @@ Lemma bor_create E κ P : ↑lftN ⊆ E → lft_ctx -∗ â–· P ={E}=∗ &{κ} P ∗ ([†κ] ={E}=∗ â–· P). Proof. - iIntros (HE) "#Hmgmt HP". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". + iIntros (HE) "#LFT HP". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iMod (ilft_create _ _ κ with "HA HI Hinv") as (A' I') "(Hκ & HA & HI & Hinv)". iDestruct "Hκ" as %Hκ. iDestruct (@big_sepS_later with "Hinv") as "Hinv". iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']". @@ -22,7 +22,8 @@ Proof. - rewrite {1}lft_inv_alive_unfold; iDestruct "Hinv" as (Pb Pi) "(Halive & Hvs & Hinh)". rewrite /lft_bor_alive; iDestruct "Halive" as (B) "(HboxB & >HownB & HB)". - iMod (lft_inh_acc _ _ P with "Hinh") as "[Hinh Hinh_close]"; first solve_ndisj. + iMod (lft_inh_extend _ _ P with "Hinh") + as "(Hinh & HIlookup & Hinh_close)"; first solve_ndisj. iMod (slice_insert_full _ _ true with "HP HboxB") as (γB) "(HBlookup & HsliceB & HboxB)"; first by solve_ndisj. rewrite lookup_fmap. iDestruct "HBlookup" as %HBlookup. @@ -45,9 +46,7 @@ Proof. iModIntro. iSplit; first by iApply lft_incl_refl. iExists γB. by iFrame. + clear -HE. iIntros "!> H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". - iAssert ⌜ is_Some (I !! κ) âŒ%I with "[#]" as %Hκ. - { iDestruct "Hinh_close" as "[H _]". by iApply "H". } - iDestruct "Hinh_close" as "[_ Hinh_close]". + iDestruct ("HIlookup" with "* HI") as %Hκ. iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hinv Hclose']". { by apply elem_of_dom. } rewrite /lft_dead; iDestruct "H†" as (Λ) "[% #H†]". @@ -71,9 +70,9 @@ Lemma bor_sep E κ P Q : ↑lftN ⊆ E → lft_ctx -∗ &{κ} (P ∗ Q) ={E}=∗ &{κ} P ∗ &{κ} Q. Proof. - iIntros (HE) "#Hmgmt HP". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". - rewrite {1}/bor /raw_bor /idx_bor_own. - iDestruct "HP" as (κ') "[#Hκκ' Htmp]". iDestruct "Htmp" as (s) "[Hbor Hslice]". + iIntros (HE) "#LFT Hbor". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". + rewrite {1}/bor. iDestruct "Hbor" as (κ') "[#Hκκ' Hbor]". + rewrite /raw_bor /idx_bor_own. iDestruct "Hbor" as (s) "[Hbor Hslice]". iDestruct (own_bor_auth with "HI Hbor") as %Hκ'. rewrite big_sepS_later big_sepS_elem_of_acc ?elem_of_dom // /lfts_inv /lft_inv /lft_inv_dead /lft_alive_in. simpl. @@ -123,11 +122,11 @@ Lemma bor_combine E κ P Q : ↑lftN ⊆ E → lft_ctx -∗ &{κ} P -∗ &{κ} Q ={E}=∗ &{κ} (P ∗ Q). Proof. - iIntros (?) "#Hmgmt HP HQ". rewrite {1 2}/bor. + iIntros (?) "#LFT HP HQ". rewrite {1 2}/bor. iDestruct "HP" as (κ1) "[#Hκ1 Hbor1]". iDestruct "HQ" as (κ2) "[#Hκ2 Hbor2]". - iMod (raw_rebor _ _ (κ1 ∪ κ2) with "Hmgmt Hbor1") as "[Hbor1 _]". + iMod (raw_rebor _ _ (κ1 ∪ κ2) with "LFT Hbor1") as "[Hbor1 _]". done. by apply gmultiset_union_subseteq_l. - iMod (raw_rebor _ _ (κ1 ∪ κ2) with "Hmgmt Hbor2") as "[Hbor2 _]". + iMod (raw_rebor _ _ (κ1 ∪ κ2) with "LFT Hbor2") as "[Hbor2 _]". done. by apply gmultiset_union_subseteq_r. iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". unfold raw_bor, idx_bor_own. iDestruct "Hbor1" as (j1) "[Hbor1 Hslice1]". iDestruct "Hbor2" as (j2) "[Hbor2 Hslice2]". diff --git a/theories/lifetime/creation.v b/theories/lifetime/creation.v index d341c5260c6a4b001ae477e63e3389266870c2c6..ff639bd742372cd0a482499abce514716150577b 100644 --- a/theories/lifetime/creation.v +++ b/theories/lifetime/creation.v @@ -97,7 +97,7 @@ Lemma raw_bor_fake' E κ P : ↑lftN ⊆ E → lft_ctx -∗ [†κ] ={E}=∗ raw_bor κ P. Proof. - iIntros (?) "#Hmgmt H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". + iIntros (?) "#LFT H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iMod (ilft_create _ _ κ with "HA HI Hinv") as (A' I') "(Hκ & HA & HI & Hinv)". iDestruct "Hκ" as %Hκ. rewrite /lft_dead. iDestruct "H†" as (Λ) "[% #H†]". iDestruct (own_alft_auth_agree A' Λ false with "HA H†") as %EQAΛ. @@ -115,7 +115,7 @@ Lemma bor_fake E κ P : ↑lftN ⊆ E → lft_ctx -∗ [†κ] ={E}=∗ &{κ}P. Proof. - iIntros (?) "#Hmgmt H†". iMod (raw_bor_fake' with "Hmgmt H†"); first done. + iIntros (?) "#LFT H†". iMod (raw_bor_fake' with "LFT H†"); first done. iModIntro. unfold bor. iExists κ. iFrame. by rewrite -lft_incl_refl. Qed. @@ -242,7 +242,7 @@ Lemma lft_create E : ↑lftN ⊆ E → lft_ctx ={E}=∗ ∃ κ, 1.[κ] ∗ â–¡ (1.[κ] ={⊤,⊤∖↑lftN}â–·=∗ [†κ]). Proof. - iIntros (?) "#Hmgmt". + iIntros (?) "#LFT". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". destruct (exist_fresh (dom (gset _) A)) as [Λ HΛ%not_elem_of_dom]. iMod (own_update with "HA") as "[HA HΛ]". diff --git a/theories/lifetime/primitive.v b/theories/lifetime/primitive.v index 5514cece8082e673483474df0681e97a63fd45a7..1b3b12419c3491fe9f7a5033584020f04af0fd66 100644 --- a/theories/lifetime/primitive.v +++ b/theories/lifetime/primitive.v @@ -355,10 +355,10 @@ Proof. Qed. (* Inheritance *) -Lemma lft_inh_acc E κ P Q : +Lemma lft_inh_extend E κ P Q : ↑inhN ⊆ E → â–· lft_inh κ false Q ={E}=∗ â–· lft_inh κ false (P ∗ Q) ∗ - (∀ I, own_ilft_auth I -∗ ⌜is_Some (I !! κ)âŒ) ∧ + (∀ I, own_ilft_auth I -∗ ⌜is_Some (I !! κ)âŒ) ∗ (∀ Q', â–· lft_inh κ true Q' ={E}=∗ ∃ Q'', â–· â–· (Q' ≡ (P ∗ Q'')) ∗ â–· P ∗ â–· lft_inh κ true Q''). Proof. @@ -371,8 +371,9 @@ Proof. iModIntro. iSplitL "Hbox HE". { iNext. rewrite /lft_inh. iExists ({[γE]} ∪ PE). rewrite to_gmap_union_singleton. iFrame. } - clear dependent PE. iSplit. - { iIntros (I) "HI". iApply (own_inh_auth with "HI HEâ—¯"). } + clear dependent PE. rewrite -(left_id_L ∅ op (â—¯ GSet {[γE]})). + iDestruct "HEâ—¯" as "[HEâ—¯' HEâ—¯]". iSplitL "HEâ—¯'". + { iIntros (I) "HI". iApply (own_inh_auth with "HI HEâ—¯'"). } iIntros (Q'). rewrite {1}/lft_inh. iDestruct 1 as (PE) "[>HE Hbox]". iDestruct (own_inh_valid_2 with "HE HEâ—¯") as %[Hle%gset_disj_included _]%auth_valid_discrete_2. diff --git a/theories/lifetime/raw_reborrow.v b/theories/lifetime/raw_reborrow.v index 95ea373a09ac547f94eb358a8540c399ca0ba6e3..c543483e23909b30f88f77b3851ce3941922bdd4 100644 --- a/theories/lifetime/raw_reborrow.v +++ b/theories/lifetime/raw_reborrow.v @@ -18,7 +18,7 @@ Lemma raw_bor_unnest E A I Pb Pi P κ i κ' : κ ⊂ κ' → lft_alive_in A κ' → Iinv -∗ idx_bor_own 1 (κ, i) -∗ slice borN i P -∗ â–· lft_bor_alive κ' Pb -∗ - â–· lft_vs κ' (idx_bor_own 1 (κ, i) ∗ (*slice borN i P ∗*) Pb) Pi ={E}=∗ ∃ Pb', + â–· lft_vs κ' (idx_bor_own 1 (κ, i) ∗ Pb) Pi ={E}=∗ ∃ Pb', Iinv ∗ raw_bor κ' P ∗ â–· lft_bor_alive κ' Pb' ∗ â–· lft_vs κ' Pb' Pi. Proof. iIntros (? Iinv Hκκ' Haliveκ') "(HI & Hκ) Hi #Hislice Hκalive' Hvs". @@ -27,8 +27,7 @@ Proof. { apply auth_update_alloc, (nat_local_update _ 0 (S n) 1); omega. } rewrite {1}/raw_bor /idx_bor_own /=. iDestruct (own_bor_auth with "HI Hi") as %?. - (* FIXME RJ: This is ugly. *) - assert (κ ⊆ κ'). { apply strict_spec_alt in Hκκ'. naive_solver. } + assert (κ ⊆ κ') by (by apply strict_include). iDestruct (lft_inv_alive_in with "Hκ") as "Hκ"; first by eauto using lft_alive_in_subseteq. rewrite lft_inv_alive_unfold; @@ -107,7 +106,7 @@ Lemma raw_rebor E κ κ' P : Proof. rewrite /lft_ctx. iIntros (??) "#LFT Hκ". destruct (decide (κ = κ')) as [<-|Hκneq]. - { iFrame. iIntros "!> #Hκ†". iMod (raw_bor_fake' with "LFT Hκ†"); done. } + { iFrame. iIntros "!> #Hκ†". by iApply (raw_bor_fake' with "LFT Hκ†"). } assert (κ ⊂ κ') by (by apply strict_spec_alt). iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". iMod (ilft_create _ _ κ' with "HA HI Hinv") as (A' I') "(% & HA & HI & Hinv)". @@ -120,32 +119,30 @@ Proof. iMod ("Hclose" with "[-Hκ]") as "_"; last auto. iNext. rewrite {2}/lfts_inv. iExists A, I. iFrame "HA HI". iApply (big_sepS_delete _ _ κ'); first by apply elem_of_dom. - iFrame "Hinv". rewrite /lft_inv /lft_inv_dead. iRight. iSplit; last done. - iExists Pi. by iFrame. } + iFrame "Hinv". rewrite /lft_inv /lft_inv_dead. iRight. + iSplit; last done. iExists Pi. by iFrame. } rewrite lft_inv_alive_unfold; iDestruct "Hκinv'" as (Pb Pi) "(Hbor & Hvs & Hinh)". rewrite {1}/raw_bor. iDestruct "Hκ" as (i) "[Hi #Hislice]". - iMod (lft_inh_acc _ _ (idx_bor_own 1 (κ, i)) with "Hinh") - as "[Hinh Hinh_close]"; first solve_ndisj. + iMod (lft_inh_extend _ _ (idx_bor_own 1 (κ, i)) with "Hinh") + as "(Hinh & HIlookup & Hinh_close)"; first solve_ndisj. iDestruct (own_bor_auth with "HI [Hi]") as %?. { by rewrite /idx_bor_own. } iDestruct (big_sepS_elem_of_acc _ _ κ with "Hinv") as "[Hκ Hκclose]". { rewrite elem_of_difference elem_of_dom not_elem_of_singleton. done. } - iMod (raw_bor_unnest _ _ _ _ (idx_bor_own 1 (κ, i) ∗ Pi)%I with "[$HI $Hκ] Hi Hislice Hbor [Hvs]") + iMod (raw_bor_unnest _ _ _ _ (idx_bor_own 1 (κ, i) ∗ Pi)%I + with "[$HI $Hκ] Hi Hislice Hbor [Hvs]") as (Pb') "([HI Hκ] & $ & Halive & Hvs)"; [solve_ndisj|done|done|..]. { iNext. by iApply lft_vs_frame. } - (* FIXME RJ: There should be sth. better than rewriting this. *) - rewrite {1}uPred.later_wand. iDestruct ("Hκclose" with "Hκ") as "Hinv". + iDestruct ("Hκclose" with "Hκ") as "Hinv". iMod ("Hclose" with "[HA HI Hinv Halive Hinh Hvs]") as "_". { iNext. rewrite {2}/lfts_inv. iExists A, I. iFrame "HA HI". iApply (big_sepS_delete _ _ κ'); first by apply elem_of_dom. iFrame "Hinv". rewrite /lft_inv. iLeft. iSplit; last done. - rewrite lft_inv_alive_unfold. iExists Pb', (idx_bor_own 1 (κ, i) ∗ Pi)%I. iFrame. } - iModIntro. iIntros "H†". - clear dependent A I Pb Pb' Pi. + rewrite lft_inv_alive_unfold. iExists Pb', (idx_bor_own 1 (κ, i) ∗ Pi)%I. + iFrame. } + clear dependent A I Pb Pb' Pi. iModIntro. iIntros "H†". iInv mgmtN as (A I) "(>HA & >HI & Hinv)" "Hclose". - iAssert ⌜is_Some (I !! κ')âŒ%I with "[#]" as %Hκ'. - { iDestruct "Hinh_close" as "[H _]". by iApply "H". } - iDestruct "Hinh_close" as "[_ Hinh_close]". + iDestruct ("HIlookup" with "* HI") as %Hκ'. iDestruct (big_sepS_delete _ _ κ' with "Hinv") as "[Hκinv' Hinv]"; first by apply elem_of_dom. rewrite {1}/lft_inv; iDestruct "Hκinv'" as "[[Halive >%]|[Hdead >%]]". @@ -160,7 +157,6 @@ Proof. iApply (big_sepS_delete _ _ κ'); first by apply elem_of_dom. iFrame "Hinv". rewrite /lft_inv /lft_inv_dead. iRight. iSplit; last done. iExists Pi'. iFrame. } - iModIntro. rewrite /raw_bor. iExists i. iFrame "∗#". + iModIntro. rewrite /raw_bor. iExists i. by iFrame. Qed. - End rebor.