diff --git a/theories/lang/heap.v b/theories/lang/heap.v index 519b83c217f569987e516c7383e119f453c187d1..679e3a9bdf438deacae375cc11a6dbb42c8eeb3b 100644 --- a/theories/lang/heap.v +++ b/theories/lang/heap.v @@ -186,7 +186,7 @@ Section heap. - by iIntros "[% [$ Hl2]]"; subst. Qed. - Lemma heap_mapsto_vec_prop_op l q1 q2 n (Φ : list val → iProp Σ) : + Lemma heap_mapsto_pred_op l q1 q2 n (Φ : list val → iProp Σ) : (∀ vl, Φ vl -∗ ⌜length vl = nâŒ) → l ↦∗{q1}: Φ ∗ l ↦∗{q2}: (λ vl, ⌜length vl = nâŒ) ⊣⊢ l ↦∗{q1+q2}: Φ. Proof. @@ -201,6 +201,13 @@ Section heap. iSplitL "Hmt1 Hown"; iExists _; by iFrame. Qed. + Lemma heap_mapsto_pred_wand l q Φ1 Φ2 : + l ↦∗{q}: Φ1 -∗ (∀ vl, Φ1 vl -∗ Φ2 vl) -∗ l ↦∗{q}: Φ2. + Proof. + iIntros "Hm Hwand". iDestruct "Hm" as (vl) "[Hm HΦ]". iExists vl. + iFrame "Hm". iApply "Hwand". done. + Qed. + Lemma heap_mapsto_vec_combine l q vl : vl ≠[] → l ↦∗{q} vl ⊣⊢ own heap_name (â—¯ [â‹… list] i ↦ v ∈ vl, @@ -210,7 +217,7 @@ Section heap. by rewrite (big_opL_commute (Auth None)) big_opL_commute1 //. Qed. - Global Instance heap_mapsto_vec_pred_fractional l (P : list val → iProp Σ): + Global Instance heap_mapsto_pred_fractional l (P : list val → iProp Σ): (∀ vl, PersistentP (P vl)) → Fractional (λ q, l ↦∗{q}: P)%I. Proof. intros p q q'. iSplit. @@ -233,7 +240,7 @@ Section heap. { rewrite -Heq /ll. done. } rewrite drop_ge; first by rewrite app_nil_r. by rewrite -Heq. Qed. - Global Instance heap_mapsto_vec_pred_as_fractional l q (P : list val → iProp Σ): + Global Instance heap_mapsto_pred_as_fractional l q (P : list val → iProp Σ): (∀ vl, PersistentP (P vl)) → AsFractional (l ↦∗{q}: P) (λ q, l ↦∗{q}: P)%I q. Proof. split. done. apply _. Qed. diff --git a/theories/typing/own.v b/theories/typing/own.v index 6bef64032d5e73f3907d3d9a61db7e7e29c8ff0e..9d99636c122de4f78a68b0ce0c1b2cdaace2336a 100644 --- a/theories/typing/own.v +++ b/theories/typing/own.v @@ -4,7 +4,7 @@ From iris.base_logic Require Import big_op. From lrust.lifetime Require Import borrow frac_borrow. From lrust.lang Require Export new_delete. From lrust.typing Require Export type. -From lrust.typing Require Import perm typing uniq_bor type_context. +From lrust.typing Require Import uniq_bor type_context. Section own. Context `{typeG Σ}. @@ -116,6 +116,24 @@ Section own. Proper (eqtype E L ==> eqtype E L) (own n). Proof. intros ?? Heq. split; f_equiv; apply Heq. Qed. + Global Instance own_send n ty : + Send ty → Send (own n ty). + Proof. + iIntros (Hsend tid1 tid2 vl) "H". iDestruct "H" as (l) "[% [Hm H†]]". subst vl. + iExists _. iSplit; first done. iFrame "H†". iNext. + iApply (heap_mapsto_pred_wand with "Hm"). iIntros (vl) "?". by iApply Hsend. + Qed. + + Global Instance own_sync n ty : + Sync ty → Sync (own n ty). + Proof. + iIntros (Hsync κ tid1 tid2 l) "H". iDestruct "H" as (l') "[Hm #Hshr]". + iExists _. iFrame "Hm". iAlways. iIntros (F q) "% Htok". + iMod ("Hshr" with "* [] Htok") as "Hfin"; first done. iModIntro. iNext. + iClear "Hshr". (* FIXME: Using "{HShr} [HShr $]" for the intro pattern in the following line doesn't work. *) + iMod "Hfin" as "[Hshr $]". by iApply Hsync. + Qed. + (** Typing *) Lemma tctx_borrow E L p n ty κ : tctx_incl E L [TCtx_hasty p (own n ty)] diff --git a/theories/typing/shr_bor.v b/theories/typing/shr_bor.v index 94e90cdcd8a9360b9c16696784cbd31f1c941d98..2f8af2758859db5a76c41943ccc69a46c8efd491 100644 --- a/theories/typing/shr_bor.v +++ b/theories/typing/shr_bor.v @@ -30,6 +30,13 @@ Section shr_bor. Global Instance subtype_shr_bor_proper E L κ : Proper (eqtype E L ==> eqtype E L) (shr_bor κ). Proof. intros ??[]. by split; apply subtype_shr_bor_mono. Qed. + + Global Instance shr_send κ ty : + Sync ty → Send (shr_bor κ ty). + Proof. + iIntros (Hsync tid1 tid2 vl) "H". iDestruct "H" as (l) "[% Hshr]". + iExists _. iSplit; first done. by iApply Hsync. + Qed. End shr_bor. Notation "&shr{ κ } ty" := (shr_bor κ ty) diff --git a/theories/typing/sum.v b/theories/typing/sum.v index 3e78528ea3b335a906d28c652ffa56c82c43479d..5288a0610f4b8cf61a330673c17b6b9a6165146e 100644 --- a/theories/typing/sum.v +++ b/theories/typing/sum.v @@ -170,7 +170,7 @@ Section sum. { rewrite <-HF. simpl. rewrite <-union_subseteq_r. apply shr_locsE_subseteq. omega. } destruct (Qp_lower_bound q'1 q'2) as (q' & q'01 & q'02 & -> & ->). - rewrite -(heap_mapsto_vec_prop_op _ q' q'02); last (by intros; apply ty_size_eq). + rewrite -(heap_mapsto_pred_op _ q' q'02); last (by intros; apply ty_size_eq). rewrite (fractional (Φ := λ q, _ ↦{q} _ ∗ _ ↦∗{q}: _)%I). iDestruct "HownC" as "[HownC1 HownC2]". iDestruct "Hown" as "[Hown1 Hown2]". iExists q'. iModIntro. iSplitL "Hown1 HownC1". diff --git a/theories/typing/uniq_bor.v b/theories/typing/uniq_bor.v index be386813b0077dc44af4631874a66756261390ec..510c204e0aa9b7225665b7268c6cd4301b6986d3 100644 --- a/theories/typing/uniq_bor.v +++ b/theories/typing/uniq_bor.v @@ -3,7 +3,7 @@ From iris.base_logic Require Import big_op. From lrust.lifetime Require Import borrow reborrow frac_borrow. From lrust.lang Require Import heap. From lrust.typing Require Export type. -From lrust.typing Require Import lft_contexts type_context shr_bor perm typing. +From lrust.typing Require Import lft_contexts type_context shr_bor. Section uniq_bor. Context `{typeG Σ}. @@ -103,6 +103,27 @@ Section uniq_bor. Global Instance subtype_uniq_proper E L κ : Proper (eqtype E L ==> eqtype E L) (uniq_bor κ). Proof. split; by apply subtype_uniq_mono. Qed. + + Global Instance uniq_send κ ty : + Send ty → Send (uniq_bor κ ty). + Proof. + iIntros (Hsend tid1 tid2 vl) "H". iDestruct "H" as (l P) "[[% #HPiff] H]". + iExists _, _. iFrame "H". iSplit; first done. iAlways. iSplit. + - iIntros "HP". iApply (heap_mapsto_pred_wand with "[HP]"). + { by iApply "HPiff". } + clear dependent vl. iIntros (vl) "?". by iApply Hsend. + - iIntros "HP". iApply "HPiff". iApply (heap_mapsto_pred_wand with "HP"). + clear dependent vl. iIntros (vl) "?". by iApply Hsend. + Qed. + + Global Instance uniq_sync κ ty : + Sync ty → Sync (uniq_bor κ ty). + Proof. + iIntros (Hsync κ' tid1 tid2 l) "H". iDestruct "H" as (l') "[Hm #Hshr]". + iExists l'. iFrame "Hm". iAlways. iIntros (F q) "% Htok". + iMod ("Hshr" with "* [] Htok") as "Hfin"; first done. iClear "Hshr". + iModIntro. iNext. iMod "Hfin" as "[Hshr $]". iApply Hsync. done. + Qed. End uniq_bor. Notation "&uniq{ κ } ty" := (uniq_bor κ ty)