diff --git a/theories/lifetime/model/accessors.v b/theories/lifetime/model/accessors.v index 3fa9066e3cd61147e94ee3d890d5a62b4d57925f..0bd9c9b71f1bc1caf51bbffb36725ff012a59fb9 100644 --- a/theories/lifetime/model/accessors.v +++ b/theories/lifetime/model/accessors.v @@ -59,7 +59,7 @@ Lemma add_vs Pb Pb' P Q Pi κ : lft_vs κ (Q ∗ Pb') Pi. Proof. iIntros "#HEQ Hvs HvsQ". iApply (lft_vs_cons with "[-Hvs] Hvs"). - iIntros "$ [HQ HPb'] #H†". + iIntros "[HQ HPb'] #H†". iApply fupd_mask_mono; last iMod ("HvsQ" with "HQ H†") as "HP". set_solver. iModIntro. iNext. iRewrite "HEQ". iFrame. Qed. diff --git a/theories/lifetime/model/borrow_sep.v b/theories/lifetime/model/borrow_sep.v index aa0c17632e6bd3a9cb23e03efdc6698412548454..b19141997e06d5c9756af415d28cd0c03cfe4c08 100644 --- a/theories/lifetime/model/borrow_sep.v +++ b/theories/lifetime/model/borrow_sep.v @@ -30,7 +30,7 @@ Proof. as (s' Pb') "(% & #HPbPb' & Hslice & Hbox)"; first solve_ndisj. { by rewrite lookup_fmap EQB. } iAssert (â–· lft_vs κ' Pb' Pi)%I with "[Hvs]" as "Hvs". - { iNext. iApply (lft_vs_cons with "[] Hvs"). iIntros "$ ? _ !>". + { iNext. iApply (lft_vs_cons with "[] Hvs"). iIntros "? _ !>". by iApply "HPbPb'". } iMod (slice_split _ _ true with "Hslice Hbox") as (γ1 γ2) "(Hγ1 & Hγ2 & % & Hs1 & Hs2 & Hbox)"; first solve_ndisj. diff --git a/theories/lifetime/model/creation.v b/theories/lifetime/model/creation.v index e0862b40427ee838eccd6bd67245cfc9d35eceb9..a9b1c3f6d1a62401fa83e52f46620cda8944d384 100644 --- a/theories/lifetime/model/creation.v +++ b/theories/lifetime/model/creation.v @@ -40,18 +40,19 @@ Proof. as "[Halive Halive']". { intros κ'. rewrite elem_of_dom. eauto. } iApply fupd_trans. iApply fupd_mask_mono; first by apply union_subseteq_l. - iMod ("Hvs" $! I with "[HI Halive Hbox Hbor] HP Hκ") as "(Hinv & HQ & Hcnt')". - { rewrite lft_vs_inv_unfold. iFrame. rewrite /lft_bor_dead. - iExists (dom _ B), P. rewrite !gset_to_gmap_dom -map_fmap_compose. - rewrite (map_fmap_ext _ ((1%Qp,.) ∘ to_agree) B); last naive_solver. - iFrame. } - rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(?&HI&Halive)". + iMod ("Hvs" $! I with "[HI Halive] HP Hκ") as "(Hinv & HQ & Hcnt')". + { rewrite lft_vs_inv_unfold. iFrame. } + rewrite lft_vs_inv_unfold; iDestruct "Hinv" as "(HI&Halive)". iSpecialize ("Halive'" with "Halive"). iMod (own_cnt_update_2 with "Hcnt Hcnt'") as "?". { apply auth_update_dealloc, (nat_local_update _ _ 0 0); lia. } rewrite /Iinv. iFrame "Hdead Halive' HI". iModIntro. iMod (lft_inh_kill with "[$Hinh $HQ]"); first set_solver+. - iModIntro. rewrite /lft_inv_dead. iExists Q. by iFrame. + iModIntro. rewrite /lft_inv_dead. iExists Q. iFrame. + rewrite /lft_bor_dead. iExists (dom _ B), P. + rewrite !gset_to_gmap_dom -map_fmap_compose. + rewrite (map_fmap_ext _ ((1%Qp,.) ∘ to_agree) B); last naive_solver. + iFrame. Qed. Lemma lfts_kill (A : gmap atomic_lft _) (I : gmap lft lft_names) (K K' : gset lft) : diff --git a/theories/lifetime/model/definitions.v b/theories/lifetime/model/definitions.v index 05cb652b8cf40335d605fa862826b590a22c9098..35d886eebbdced1f241d0ea842f53b6030b2a19a 100644 --- a/theories/lifetime/model/definitions.v +++ b/theories/lifetime/model/definitions.v @@ -131,8 +131,7 @@ Section defs. Definition lft_vs_inv_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → iProp Σ) (I : gmap lft lft_names) : iProp Σ := - (lft_bor_dead κ ∗ - own_ilft_auth I ∗ + (own_ilft_auth I ∗ [∗ set] κ' ∈ dom _ I, ∀ Hκ : κ' ⊂ κ, lft_inv_alive κ' Hκ)%I. Definition lft_vs_go (κ : lft) (lft_inv_alive : ∀ κ', κ' ⊂ κ → iProp Σ) @@ -230,7 +229,7 @@ Lemma lft_vs_inv_go_ne κ (f f' : ∀ κ', κ' ⊂ κ → iProp Σ) I n : (∀ κ' (Hκ : κ' ⊂ κ), f κ' Hκ ≡{n}≡ f' κ' Hκ) → lft_vs_inv_go κ f I ≡{n}≡ lft_vs_inv_go κ f' I. Proof. - intros Hf. apply sep_ne, sep_ne, big_opS_ne=> // κ' _. + intros Hf. apply sep_ne, big_opS_ne=> // κ' _. by apply forall_ne=> Hκ. Qed. @@ -259,7 +258,7 @@ Proof. apply (Fix_F_proper _ (λ _, dist n)); auto using lft_inv_alive_go_ne. Qed. Lemma lft_vs_inv_unfold κ (I : gmap lft lft_names) : - lft_vs_inv κ I ⊣⊢ lft_bor_dead κ ∗ + lft_vs_inv κ I ⊣⊢ own_ilft_auth I ∗ [∗ set] κ' ∈ dom _ I, ⌜κ' ⊂ κ⌠→ lft_inv_alive κ'. Proof. diff --git a/theories/lifetime/model/primitive.v b/theories/lifetime/model/primitive.v index 27368e2108de48c974e5cdfcb0fcd9afd71a2cdd..1d1d3a5537e00548f07ad1509786332bb5590a59 100644 --- a/theories/lifetime/model/primitive.v +++ b/theories/lifetime/model/primitive.v @@ -490,15 +490,15 @@ Proof. Qed. Lemma lft_vs_cons κ Pb Pb' Pi : - (lft_bor_dead κ -∗ â–· Pb'-∗ [†κ] ={↑borN}=∗ lft_bor_dead κ ∗ â–· Pb) -∗ + (â–· Pb'-∗ [†κ] ={↑borN}=∗ â–· Pb) -∗ lft_vs κ Pb Pi -∗ lft_vs κ Pb' Pi. Proof. iIntros "Hcons Hvs". rewrite !lft_vs_unfold. iDestruct "Hvs" as (n) "[Hnâ— Hvs]". iExists n. iFrame "Hnâ—". iIntros (I). rewrite {1}lft_vs_inv_unfold. - iIntros "(Hdead & Hinv & Hκs) HPb #Hκ†". - iMod ("Hcons" with "Hdead HPb Hκ†") as "[Hdead HPb]". - iApply ("Hvs" $! I with "[Hdead Hinv Hκs] HPb Hκ†"). + iIntros "(Hinv & Hκs) HPb #Hκ†". + iMod ("Hcons" with "HPb Hκ†") as "HPb". + iApply ("Hvs" $! I with "[Hinv Hκs] HPb Hκ†"). rewrite lft_vs_inv_unfold. by iFrame. Qed. End primitive. diff --git a/theories/lifetime/model/reborrow.v b/theories/lifetime/model/reborrow.v index 283317ab260b58152abda4b70494606ce3e14353..a6ad7bb0f1c4f3b9c0cd8e236af97c4637b63e0a 100644 --- a/theories/lifetime/model/reborrow.v +++ b/theories/lifetime/model/reborrow.v @@ -68,7 +68,7 @@ Proof. clear dependent Iinv I. iNext. rewrite lft_vs_unfold. iExists (S n). iFrame "Hnâ—". iIntros (I) "Hinv [HP HPb] #Hκ†". - rewrite {1}lft_vs_inv_unfold; iDestruct "Hinv" as "(Hκdead' & HI & Hinv)". + rewrite {1}lft_vs_inv_unfold; iDestruct "Hinv" as "(HI & Hinv)". iDestruct (own_bor_auth with "HI Hi") as %?%(elem_of_dom (D:=gset lft)). iDestruct (@big_sepS_delete with "Hinv") as "[Hκalive Hinv]"; first done. rewrite lft_inv_alive_unfold. @@ -85,9 +85,9 @@ Proof. { by rewrite lookup_fmap HB. } iDestruct (@big_sepM_delete with "HB") as "[Hcnt HB]"; first done. rewrite /=. iDestruct "Hcnt" as "[% H1â—¯]". - iMod ("Hvs" $! I with "[Hκdead' HI Hinv Hvs' Hinh HBâ— Hbox HB] + iMod ("Hvs" $! I with "[HI Hinv Hvs' Hinh HBâ— Hbox HB] [$HPb $Hi] Hκ†") as "($ & $ & Hcnt')". - { rewrite lft_vs_inv_unfold. iFrame "Hκdead' HI". + { rewrite lft_vs_inv_unfold. iFrame "HI". iApply (big_sepS_delete _ (dom (gset lft) I) with "[- $Hinv]"); first done. iIntros (_). rewrite lft_inv_alive_unfold. iExists Pb', Pi'. iFrame "Hvs' Hinh". rewrite /lft_bor_alive. @@ -135,7 +135,7 @@ Proof. { rewrite /lft_bor_alive (big_sepM_delete _ B i') //. iDestruct "HB" as "[_ HB]". iExists (delete i' B). rewrite -fmap_delete. iFrame. rewrite fmap_delete -insert_delete delete_insert ?lookup_delete //=. } - { iNext. iApply (lft_vs_cons with "[] Hvs"). iIntros "$ [??] _ !>". iNext. + { iNext. iApply (lft_vs_cons with "[] Hvs"). iIntros "[??] _ !>". iNext. iRewrite "HeqPb'". iFrame. by iApply "HP'". } iDestruct "HH" as "([HI Hinvκ] & $ & Halive & Hvs)". iApply "Hclose". iExists _, _. iFrame.