diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 0a1248a74245c07abcf33c1ac2380321efd39aae..11c498bcd9e20d10594e75f0d081d2e0c8376ae0 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -93,23 +93,21 @@ Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := (l ↦ NONEV ∗ tok γt)%I ∨ - (∃ v m, l ↦ SOMEV v ∗ + (∃ v m p, l ↦ SOMEV v ∗ iProto_own γ i (<(Send, j)> m)%proto ∗ - (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ + iMsg_car m v (Next p) ∗ own γE (â—E (Next p))) ∨ (∃ p, l ↦ NONEV ∗ iProto_own γ i p ∗ own γE (â—E (Next p))). Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', + ∃ γ (γEs : list gname) (l:loc) (i:nat) (n:nat) p', ⌜ c = (#l,#n,#i)%V ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ - ([∗list] j ↦ _ ∈ replicate n (), - ∃ γE2 γt1 γt2, - inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ - inv (nroot.@"p") (chan_inv γ γE2 γt2 j i (l +â‚— (pos n j i)))) ∗ + (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ + ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ â–· (p' ⊑ p) ∗ - own γE1 (â—E (Next p')) ∗ own γE1 (â—¯E (Next p')) ∗ + own (γEs !!! i) (â—E (Next p')) ∗ own (γEs !!! i) (â—¯E (Next p')) ∗ iProto_own γ i p'. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). @@ -125,15 +123,13 @@ Definition chan_pool `{!heapGS Σ, !chanG Σ} ∃ γ (γEs : list gname) (n:nat) (l:loc), ⌜cs = (#l,#n)%V⌠∗ ⌜(i' + length ps = n)%nat⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ + (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ + ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ [∗ list] i ↦ _ ∈ replicate n (), (∀ p, ⌜i' <= i⌠-∗ ⌜ps !! (i - i') = Some p⌠-∗ own (γEs !!! i) (â—E (Next p)) ∗ own (γEs !!! i) (â—¯E (Next p)) ∗ - iProto_own γ i p) ∗ - [∗ list] j ↦ _ ∈ replicate n (), - ∃ γt1 γt2, - inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt1 i j (l +â‚— (pos n i j))) ∗ - inv (nroot.@"p") (chan_inv γ (γEs !!! j) γt2 j i (l +â‚— (pos n j i))). + iProto_own γ i p). Section channel. Context `{!heapGS Σ, !chanG Σ}. @@ -269,30 +265,17 @@ Section channel. iSplit. { done. } iFrame "IHp". + iSplit. + { iIntros (i j Hi Hj). + rewrite (big_sepL_lookup _ _ i); [|by rewrite lookup_replicate]. + rewrite (big_sepL_lookup _ _ j); [|by rewrite lookup_replicate]. + iApply "IH". } iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - iSplitL. - { iIntros (p Hle' HSome''). - iFrame. rewrite right_id_L in HSome''. - rewrite (list_lookup_total_alt ps). - rewrite HSome''. simpl. iFrame. } - iApply big_sepL_intro. - iIntros "!>" (j ? HSome''). - assert (i < n) as Hle'. - { apply lookup_replicate in HSome' as [_ Hle']. done. } - assert (j < n) as Hle''. - { apply lookup_replicate in HSome'' as [_ Hle'']. done. } - iDestruct (big_sepL_lookup _ _ i () with "IH") as "IH'"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ j () with "IH'") as "IH''"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ j () with "IH") as "H"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ i () with "H") as "H'"; - [by rewrite lookup_replicate|]. - iDestruct "IH''" as (γ1) "?". - iDestruct "H'" as (γ2) "?". - iExists _, _. iFrame "#". + iIntros (p Hle' HSome''). + iFrame. rewrite right_id_L in HSome''. + rewrite (list_lookup_total_alt ps). + rewrite HSome''. simpl. iFrame. Qed. Lemma get_chan_spec cs i ps p : @@ -301,11 +284,10 @@ Section channel. {{{ c, RET c; c ↣ p ∗ chan_pool cs (i+1) ps }}}. Proof. iIntros (Φ) "Hcs HΦ". - iDestruct "Hcs" as (γp γEs n l -> <-) "[#IHp Hl]". - wp_lam. wp_pures. - simpl. + iDestruct "Hcs" as (γp γEs n l -> <-) "(#IHp & #IHl & Hl)". + wp_lam. wp_pures. rewrite replicate_add. simpl. - iDestruct "Hl" as "[Hl1 [[Hi #IHs] Hl3]]". simpl. + iDestruct "Hl" as "[Hl1 [Hi Hl3]]". iDestruct ("Hi" with "[] []") as "(Hauth & Hown & Hp)". { rewrite right_id_L. rewrite replicate_length. iPureIntro. lia. } { rewrite right_id_L. rewrite replicate_length. @@ -316,16 +298,11 @@ Section channel. { rewrite iProto_pointsto_eq. iExists _, _, _, _, _, _. iSplit; [done|]. rewrite replicate_length. rewrite right_id_L. - iFrame "#∗". - iSplit; [|iNext; done]. - rewrite replicate_add. - iApply (big_sepL_impl with "IHs"). - iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[??]". - iExists _,_,_. iFrame. } - iExists _, _, _, _. iSplit; [done|]. + iFrame. iFrame "#∗". iNext. done. } + iExists γp, γEs, _, _. iSplit; [done|]. iSplit. { iPureIntro. lia. } - iFrame "#∗". + iFrame. iFrame "#∗". rewrite replicate_add. simpl. iSplitL "Hl1". @@ -333,7 +310,7 @@ Section channel. iIntros "!>" (i' ? HSome''). assert (i' < i). { rewrite lookup_replicate in HSome''. lia. } - iIntros "[H $]" (p' Hle'). lia. } + iIntros "H" (p' Hle'). lia. } simpl. iFrame "#∗". iSplitR. @@ -343,7 +320,7 @@ Section channel. iIntros "!>" (i' ? HSome''). assert (i' < length ps). { rewrite lookup_replicate in HSome''. lia. } - iIntros "[H $]" (p' Hle' HSome). + iIntros "H" (p' Hle' HSome). iApply "H". { iPureIntro. lia. } iPureIntro. @@ -373,44 +350,42 @@ Section channel. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". wp_bind (Fst _). - iInv "IH" as "HI". + iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". + iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hi [Hown Hctx]]". + { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. + iFrame. done. } iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ n))%I with "[HI Hown]" - as "[HI [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$ $]]". + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hj [Hown Hctx]]". + { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$ $]]". iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". - iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; - [by apply lookup_replicate_2|]. - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hi" as %Hi. + iDestruct "Hj" as %Hj. + iDestruct ("Hls" $! i j with "[//] [//]") as (γt) "IHl1". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". + - iDestruct "HIp" as (? m p'') "(>Hl & Hown' & HIp)". wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p'') "[Hle' Hown]". - iDestruct "Hown'" as (p''') "[Hle'' Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. + iDestruct (iProto_own_excl with "Hown Hown'") as %[]. - iDestruct "HIp" as (p'') "(>Hl' & Hown' & HIp)". wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p''') "[Hle' Hown]". - iDestruct "Hown'" as (p'''') "[Hle'' Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } + iDestruct (iProto_own_excl with "Hown Hown'") as %[]. } iDestruct "HIp" as "[>Hl' Htok]". wp_store. iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]"; [by apply excl_auth_update|]. iModIntro. iSplitL "Hl' Hâ— Hown Hle". - { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + { iRight. iLeft. iIntros "!>". iExists _, _, _. iFrame. iSplitL "Hown Hle"; [by iApply (iProto_own_le with "Hown Hle")|]. - iExists _. iFrame. by rewrite iMsg_base_eq. } + by rewrite iMsg_base_eq. } wp_pures. iLöb as "HL". wp_lam. @@ -420,10 +395,10 @@ Section channel. { iDestruct "HIp" as ">[Hl' Htok']". iDestruct (own_valid_2 with "Htok Htok'") as %[]. } iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl' & Hown & HIp)". + - iDestruct "HIp" as (? m p'') "(>Hl' & Hown & HIp)". wp_load. iModIntro. iSplitL "Hl' Hown HIp". - { iRight. iLeft. iExists _, _. iFrame. } + { iRight. iLeft. iExists _, _,_. iFrame. } wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). - iDestruct "HIp" as (p'') "(>Hl' & Hown & Hâ—)". wp_load. @@ -466,19 +441,23 @@ Section channel. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". do 6 wp_pure _. wp_bind (Fst _). wp_pure _. - iInv "IH" as "HI". + iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". + iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hi [Hown Hctx]]". + { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. + iFrame. done. } iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ - iProto_ctx γ n))%I with "[HI Hown]" as "[HI [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$$]]". + iProto_ctx γ n))%I with "[Hctx Hown]" as "[Hj [Hown Hctx]]". + { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$$]]". iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". - iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; - [by apply lookup_replicate_2|]. - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hi" as %Hi. + iDestruct "Hj" as %Hj. + iDestruct ("Hls" $! j i with "[//] [//]") as (γt) "IHl2". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". @@ -497,8 +476,7 @@ Section channel. { iRight. iRight. iExists _. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } - iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". - iDestruct "HIp" as (p'') "[Hm Hp']". + iDestruct "HIp" as (w m p'') "(>Hl' & Hown' & Hm & Hp')". iInv "IH" as "Hctx". wp_xchg. iDestruct (iProto_own_le with "Hown Hle") as "Hown". @@ -519,28 +497,4 @@ Section channel. iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. - Lemma iProto_le_select_l {TT1 TT2:tele} j - (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : - ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ - (<(Send,j) @.. (tt:TT1)> MSG (InjLV (v1 tt)) {{ P1 tt }} ; p1 tt). - Proof. - rewrite /iProto_choice. - iApply iProto_le_trans; last first. - { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. - iApply iProto_le_refl. } - iIntros (tt). by iExists (inl tt). - Qed. - - Lemma iProto_le_select_r {TT1 TT2:tele} j - (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : - ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ - (<(Send,j) @.. (tt:TT2)> MSG (InjRV (v2 tt)) {{ P2 tt }} ; p2 tt). - Proof. - rewrite /iProto_choice. - iApply iProto_le_trans; last first. - { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. - iApply iProto_le_refl. } - iIntros (tt). by iExists (inr tt). - Qed. - End channel. diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 07c044920f750bd14e942c5043f7cc2e6ba20891..a125a6f8a15c6820907d0f495eb746a972770a45 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -344,7 +344,7 @@ Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ") wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; eexists x6; eexists x7; eexists x8) with pat. -Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : +Lemma iProto_consistent_equiv_proof {Σ} (ps : list (iProto Σ)) : (∀ i j, valid_target ps i j) ∗ (∀ i j m1 m2, (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ @@ -396,56 +396,32 @@ Tactic Notation "iProto_consistent_take_step_step" := let m1 := fresh in let m2 := fresh in iIntros (i j m1 m2) "#Hm1 #Hm2"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; + repeat (destruct i as [|i]=> /=; + [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq=> /=| + try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]); + try (rewrite lookup_total_nil); + try (by rewrite iProto_end_message_equivI); iDestruct (iProto_message_equivI with "Hm2") - as "[%Heq2 Hm2']";simplify_eq; + as "[%Heq2 Hm2']";simplify_eq=> /=; try (iClear "Hm1' Hm2'"; iExists _,_,_,_,_,_,_,_,_,_; iSplitL "Hm1"; [iFrame "#"|]; iSplitL "Hm2"; [iFrame "#"|]; iSplit; [iPureIntro; tc_solve|]; iSplit; [iPureIntro; tc_solve|]; - simpl; iClear "Hm1 Hm2"; clear m1 m2); - try (repeat (rewrite (insert_commute _ _ i); [|done]); - rewrite insert_insert; - repeat (rewrite (insert_commute _ _ j); [|done]); - rewrite insert_insert). + simpl; iClear "Hm1 Hm2"; clear m1 m2). Tactic Notation "iProto_consistent_take_step_target" := let i := fresh in iIntros (i j a m); rewrite /valid_target; - iIntros "#Hm"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; - iDestruct (iProto_message_equivI with "Hm") - as "[%Heq Hm']";simplify_eq; - repeat (try rewrite lookup_empty; - try rewrite lookup_insert; - rewrite lookup_insert_ne; [|lia]); - try rewrite lookup_insert; try done. + iIntros "#Hm1"; + repeat (destruct i as [|i]=> /=; + [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + by iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']" ; simplify_eq=> /=| + try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]). Tactic Notation "iProto_consistent_take_step" := try iNext; @@ -463,3 +439,6 @@ Tactic Notation "iProto_consistent_resolve_step" := Tactic Notation "iProto_consistent_take_steps" := repeat (iProto_consistent_take_step; iProto_consistent_resolve_step). + +Tactic Notation "wp_get_chan" "(" simple_intropattern(c) ")" constr(pat) := + wp_smart_apply (get_chan_spec with "[$]"); iIntros (c); iIntros pat. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index e5a2b781b7c649a66ea6594969299667b21f4f3a..aa0826a901447c221041396c7a8b895ac871a5e1 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -499,14 +499,14 @@ Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate EN Definition can_step {Σ V} (rec : list (iProto Σ V) → iProp Σ) (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, - (ps !!! i ≡ <(Send, j)> m1) -∗ - (ps !!! j ≡ <(Recv, i)> m2) -∗ + (ps !!! i ≡ (<(Send, j)> m1)) -∗ + (ps !!! j ≡ (<(Recv, i)> m2)) -∗ ∀ v p1, iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). Definition valid_target {Σ V} (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := - ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. + ∀ a m, (ps !!! i ≡ (<(a, j)> m)) -∗ ⌜is_Some (ps !! j)âŒ. Definition iProto_consistent_pre {Σ V} (rec : list (iProto Σ V) → iProp Σ) (ps : list (iProto Σ V)) : iProp Σ := @@ -600,8 +600,7 @@ Definition iProto_own_frag `{!protoG Σ V} (γ : gname) Definition iProto_own_auth `{!protoG Σ V} (γ : gname) (ps : list (iProto Σ V)) : iProp Σ := - own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> - (list_to_map (zip (seq 0 (length ps)) ps))) : gmap _ _)). + own γ (gmap_view_auth (DfracOwn 1) ((λ p, Excl' (Next p)) <$> map_seq 0 ps)). Definition iProto_ctx `{protoG Σ V} (γ : gname) (ps_len : nat) : iProp Σ := @@ -781,44 +780,191 @@ Section proto. iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. iDestruct "H" as (->) "H". by iExists _. Qed. + + Lemma iProto_le_send i m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_recv i m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_base a v P p1 p2 : + â–· (p1 ⊑ p2) -∗ + (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". + iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". + iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + Qed. + + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). + destruct (iProto_case p3) as [->|([]&i&m3&->)]. + - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". + iRewrite "Hp2" in "H1"; clear p2. + iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. + iApply iProto_le_send. iIntros (v p3') "Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". + iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). + - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". + iRewrite "Hp2" in "H1". + iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". + iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". + iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". + iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). + Qed. + + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. + - iApply iProto_le_end. + - iApply iProto_le_send. auto 10 with iFrame. + - iApply iProto_le_recv. auto 10 with iFrame. + Qed. + + + Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). + Proof. solve_proper. Qed. + + Lemma iProto_own_auth_agree γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !! i ≡ Some p). + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + rewrite gmap_view_both_validI. + iDestruct "H✓" as "[_ [H1 H2]]". + rewrite lookup_fmap. + simpl. + rewrite lookup_map_seq_0. + destruct (ps !! i) eqn:Heqn; last first. + { rewrite Heqn. rewrite !option_equivI. done. } + rewrite Heqn. + simpl. rewrite !option_equivI excl_equivI. by iNext. + Qed. + + Lemma iProto_own_auth_agree_Some γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ ⌜is_Some (ps !! i)âŒ. + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + rewrite gmap_view_both_validI. + iDestruct "H✓" as "[_ [H1 H2]]". + rewrite lookup_fmap. + simpl. + rewrite lookup_map_seq_0. + destruct (ps !! i) eqn:Heqn; last first. + { rewrite Heqn. rewrite !option_equivI. done. } + rewrite Heqn. + simpl. rewrite !option_equivI excl_equivI. done. + Qed. + + Lemma iProto_own_auth_update γ ps i p p' : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ + iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (iProto_own_auth_agree_Some with "Hâ— Hâ—¯") as %HSome. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. + { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } + iFrame. rewrite -fmap_insert. + rewrite /iProto_own_auth. + rewrite insert_map_seq_0; [done|]. + by apply lookup_lt_is_Some_1. + Qed. + + Lemma iProto_own_auth_alloc ps : + ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. + Proof. + iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". + { apply gmap_view_auth_valid. } + iExists γ. + iInduction ps as [|p ps] "IH" using rev_ind. + { iModIntro. iFrame. done. } + iMod ("IH" with "Hauth") as "[Hauth Hfrags]". + iFrame "Hfrags". + iMod (own_update with "Hauth") as "[Hauth Hfrag]". + { apply (gmap_view_alloc _ (length ps) (DfracOwn 1) (Excl' (Next p))); [|done|done]. + rewrite fmap_map_seq. + rewrite lookup_map_seq_0. + apply lookup_ge_None_2. rewrite fmap_length. done. } + simpl. + iModIntro. + rewrite right_id_L. + rewrite -fmap_insert. iFrame. + iSplitL "Hauth". + - rewrite /iProto_own_auth. + rewrite map_seq_snoc. simpl. done. + - iSplit; [|done]. + iExists _. iFrame. by iApply iProto_le_refl. + Qed. + + Lemma list_lookup_Some_le (ps : list $ iProto Σ V) (i : nat) (p1 : iProto Σ V) : + ⊢@{iProp Σ} ps !! i ≡ Some p1 -∗ ⌜i < length psâŒ. + Proof. + iIntros "HSome". + rewrite option_equivI. + destruct (ps !! i) eqn:Heqn; [|done]. + iPureIntro. + by apply lookup_lt_is_Some_1. + Qed. Lemma valid_target_le ps i p1 p2 : (∀ i' j', valid_target ps i' j') -∗ - ps !!! i ≡ p1 -∗ + ps !! i ≡ Some p1 -∗ p1 ⊑ p2 -∗ (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. - Proof. Admitted. - (* iIntros "Hprot #HSome Hle". *) - (* pose proof (iProto_case p1) as [Hend|Hmsg]. *) - (* { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". *) - (* iFrame "Hle". *) - (* iIntros (i' j' a m) "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { Search list_lookup_total insert. rewrite list_lookup_total_insert. ; [done|]. lia. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "H" in "Hm". *) - (* by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". } *) - (* destruct Hmsg as (t & n & m & Hmsg). *) - (* setoid_rewrite Hmsg. *) - (* iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". *) - (* iIntros (i' j' a m') "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { rewrite lookup_insert. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". *) - (* iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". *) - (* simplify_eq. by iApply "Hprot". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". *) - (* Qed. *) + Proof. + iIntros "Hprot #HSome Hle". + iDestruct (list_lookup_Some_le with "HSome") as %Hi. + pose proof (iProto_case p1) as [Hend|Hmsg]. + { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". + iFrame "Hle". + iIntros (i' j' a m) "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { rewrite list_lookup_insert; [done|]. done. } + rewrite (list_lookup_insert_ne _ i j'); [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite list_lookup_total_insert; [|done]. iRewrite "H" in "Hm". + by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } + rewrite list_lookup_total_insert_ne; [|done]. + by iApply "Hprot". } + destruct Hmsg as (t & n & m & Hmsg). + setoid_rewrite Hmsg. + iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". + iIntros (i' j' a m') "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { by rewrite list_lookup_insert. } + rewrite (list_lookup_insert_ne _ i j'); [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite list_lookup_total_insert; [|done]. iRewrite "Heq" in "Hm". + iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". + simplify_eq. iApply ("Hprot" $! i'). + rewrite list_lookup_total_alt. iRewrite "HSome". done. } + rewrite list_lookup_total_insert_ne; [|done]. + by iApply "Hprot". + Qed. Lemma iProto_consistent_le ps i p1 p2 : iProto_consistent ps -∗ - ps !!! i ≡ p1 -∗ + ps !! i ≡ Some p1 -∗ p1 ⊑ p2 -∗ iProto_consistent (<[i := p2]>ps). Proof. @@ -826,19 +972,21 @@ Section proto. iRevert "HSome". iLöb as "IH" forall (p1 p2 ps). iIntros "#HSome". + iDestruct (list_lookup_Some_le with "HSome") as %Hi. rewrite !iProto_consistent_unfold. iDestruct "Hprot" as "(Htar & Hprot)". iDestruct (valid_target_le with "Htar HSome Hle") as "[Htar Hle]". iFrame. iIntros (i' j' m1 m2) "#Hm1 #Hm2". destruct (decide (i = i')) as [<-|Hneq]. - { rewrite list_lookup_total_insert; [|admit]. + { rewrite list_lookup_total_insert; [|done]. pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + { setoid_rewrite Hend. + rewrite !option_equivI. rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). setoid_rewrite Hmsg. destruct a; last first. - { rewrite iProto_message_equivI. + { rewrite !option_equivI. rewrite iProto_message_equivI. iDestruct "Hm1" as "[%Htag Hm1]". done. } rewrite iProto_message_equivI. iDestruct "Hm1" as "[%Htag Hm1]". @@ -850,29 +998,35 @@ Section proto. iDestruct "Hle" as (m') "[#Heq H]". iDestruct ("H" with "Hm1'") as (p') "[Hle H]". destruct (decide (i = j')) as [<-|Hneq]. - { rewrite list_lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Heq _]". done. admit. } + { rewrite list_lookup_total_insert; [|done]. + rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Heq _]". done. } iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". - { iRewrite -"Heq". rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { iRewrite -"Heq". iEval (rewrite list_lookup_total_alt). + iRewrite "HSome". done. } { rewrite list_lookup_total_insert_ne; [|done]. done. } iDestruct "Hprot" as (p'') "[Hm Hprot]". iExists p''. iFrame. iNext. iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". - { rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_insert; [done|]. + by rewrite insert_length. } iClear "IH Hm1 Hm2 Heq". rewrite list_insert_insert. rewrite (list_insert_commute _ j' i); [|done]. rewrite list_insert_insert. done. } rewrite list_lookup_total_insert_ne; [|done]. destruct (decide (i = j')) as [<-|Hneq']. - { rewrite list_lookup_total_insert. + { rewrite list_lookup_total_insert; [|done]. pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + { setoid_rewrite Hend. + rewrite !option_equivI. + rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). setoid_rewrite Hmsg. destruct a. - { rewrite iProto_message_equivI. + { rewrite !option_equivI. + rewrite iProto_message_equivI. iDestruct "Hm2" as "[%Htag Hm2]". done. } rewrite iProto_message_equivI. iDestruct "Hm2" as "[%Htag Hm2]". @@ -882,17 +1036,17 @@ Section proto. iDestruct "Hle" as (m') "[#Heq Hle]". iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } - { rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { iEval (rewrite list_lookup_total_alt). iRewrite "HSome". done. } iDestruct ("Hprot") as (p') "[Hm1' Hprot]". iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". iSpecialize ("Hm2" $! v (Next p2')). iExists p2'. iRewrite -"Hm2". iFrame. iDestruct ("IH" with "Hprot Hle []") as "HI". - { iPureIntro. rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert. done. admit. } + { iPureIntro. rewrite list_lookup_insert_ne; [|done]. + by rewrite list_lookup_insert. } rewrite list_insert_commute; [|done]. - rewrite !list_insert_insert. done. admit. } + rewrite !list_insert_insert. done. } rewrite list_lookup_total_insert_ne; [|done]. iIntros (v p) "Hm1'". iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". @@ -902,68 +1056,9 @@ Section proto. rewrite (list_insert_commute _ j' i); [|done]. rewrite (list_insert_commute _ i' i); [|done]. iApply ("IH" with "Hprot Hle []"). - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. done. - Admitted. - - Lemma iProto_le_send i m1 m2 : - (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ - (<(Send,i)> m1) ⊑ (<(Send,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_recv i m1 m2 : - (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ - (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_base a v P p1 p2 : - â–· (p1 ⊑ p2) -∗ - (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. - - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". - iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". - iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - Qed. - - Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). - destruct (iProto_case p3) as [->|([]&i&m3&->)]. - - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". - - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". - iRewrite "Hp2" in "H1"; clear p2. - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. - iApply iProto_le_send. iIntros (v p3') "Hm3". - iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". - iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". - iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). - - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". - iRewrite "Hp2" in "H1". - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". - iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". - iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". - iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". - iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). - Qed. - - Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. - - iApply iProto_le_end. - - iApply iProto_le_send. auto 10 with iFrame. - - iApply iProto_le_recv. auto 10 with iFrame. Qed. Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. @@ -1115,16 +1210,18 @@ Section proto. Lemma iProto_consistent_target ps m a i j : iProto_consistent ps -∗ - ps !!! i ≡ (<(a, j)> m) -∗ + ps !! i ≡ Some (<(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. Proof. - rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". iApply "Htar". + rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". + iIntros "H". iApply ("Htar" $! i). + rewrite list_lookup_total_alt. iRewrite "H". done. Qed. Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ - ps !!! i ≡ (<(Send, j)> m1) -∗ - ps !!! j ≡ (<(Recv, i)> m2) -∗ + ps !! i ≡ Some (<(Send, j)> m1) -∗ + ps !! j ≡ Some (<(Recv, i)> m2) -∗ iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). @@ -1132,52 +1229,12 @@ Section proto. iIntros "Hprot #Hi #Hj Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. iDestruct "Hprot" as "[_ Hprot]". - iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". + iDestruct ("Hprot" $! i j with "[Hi] [Hj] Hm1") as (p2) "[Hm2 Hprot]". + { rewrite list_lookup_total_alt. iRewrite "Hi". done. } + { rewrite list_lookup_total_alt. iRewrite "Hj". done. } iExists p2. iFrame. Qed. - Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). - Proof. solve_proper. Qed. - - Lemma iProto_own_auth_agree γ ps i p : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). - Proof. Admitted. - (* iIntros "Hâ— Hâ—¯". *) - (* iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) - (* rewrite gmap_view_both_validI. *) - (* iDestruct "H✓" as "[_ [H1 H2]]". *) - (* rewrite list_lookup_total_alt lookup_fmap. *) - (* destruct (ps !! i); last first. *) - (* { rewrite !option_equivI. } *) - (* simpl. rewrite !option_equivI excl_equivI. by iNext. *) - (* Qed. *) - - Lemma iProto_own_auth_update γ ps i p p' : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ - iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. - Proof. - iIntros "Hâ— Hâ—¯". - iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. - { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } - iFrame. rewrite -fmap_insert. Admitted. - - Lemma iProto_own_auth_alloc ps : - ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. - Proof. Admitted. - (* iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". *) - (* { apply gmap_view_auth_valid. } *) - (* iExists γ. *) - (* iInduction ps as [|i p ps Hin] "IH" using map_ind. *) - (* { iModIntro. iFrame. by iApply big_sepM_empty. } *) - (* iMod ("IH" with "Hauth") as "[Hauth Hfrags]". *) - (* rewrite big_sepM_insert; [|done]. iFrame "Hfrags". *) - (* iMod (own_update with "Hauth") as "[Hauth Hfrag]". *) - (* { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. *) - (* by rewrite lookup_fmap Hin. } *) - (* iModIntro. rewrite -fmap_insert. iFrame. *) - (* iExists _. iFrame. iApply iProto_le_refl. *) - (* Qed. *) - Lemma iProto_own_le γ s p1 p2 : iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. Proof. @@ -1185,6 +1242,29 @@ Section proto. iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). Qed. + Lemma iProto_own_excl γ i (p1 p2 : iProto Σ V) : + iProto_own γ i p1 -∗ iProto_own γ i p2 -∗ False. + Proof. + rewrite /iProto_own. + iDestruct 1 as (p1') "[_ Hp1]". + iDestruct 1 as (p2') "[_ Hp2]". + iDestruct (own_prot_excl with "Hp1 Hp2") as %[]. + Qed. + + Lemma iProto_ctx_agree γ n i p : + iProto_ctx γ n -∗ + iProto_own γ i p -∗ + ⌜i < nâŒ. + Proof. + iIntros "Hctx Hown". + rewrite /iProto_ctx /iProto_own. + iDestruct "Hctx" as (ps <-) "[Hauth Hps]". + iDestruct "Hown" as (p') "[Hle Hown]". + iDestruct (iProto_own_auth_agree_Some with "Hauth Hown") as %HSome. + iPureIntro. + by apply lookup_lt_is_Some_1. + Qed. + Lemma iProto_init ps : â–· iProto_consistent ps -∗ |==> ∃ γ, iProto_ctx γ (length ps) ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. @@ -1194,7 +1274,6 @@ Section proto. iExists γ. iFrame. iExists _. by iFrame. Qed. - Lemma iProto_step γ ps_dom i j m1 m2 p1 v : iProto_ctx γ ps_dom -∗ iProto_own γ i (<(Send, j)> m1) -∗ @@ -1204,49 +1283,48 @@ Section proto. iProto_own γ i p1 ∗ iProto_own γ j p2. Proof. iIntros "Hctx Hi Hj Hm". + iDestruct (iProto_ctx_agree with "Hctx Hi") as %Hi. + iDestruct (iProto_ctx_agree with "Hctx Hj") as %Hij. iDestruct "Hi" as (pi) "[Hile Hi]". iDestruct "Hj" as (pj) "[Hjle Hj]". iDestruct "Hctx" as (ps Hdom) "[Hauth Hconsistent]". iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". iDestruct (own_prot_idx with "Hi Hj") as %Hneq. - iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". - { by rewrite list_lookup_total_insert_ne. } + iAssert (â–· (<[i:=<(Send, j)> m1]>ps !! j ≡ Some pj))%I as "Hpj'". + { by rewrite list_lookup_insert_ne. } iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" as "[Hi' Hile]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpi". rewrite list_lookup_total_alt. - destruct (ps !! i); [done|]. - iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } + iFrame. iRewrite "Heq" in "Hpi". destruct (ps !! i); [done|]. + by rewrite option_equivI. } iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" as "[Hj' Hjle]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpj". rewrite !list_lookup_total_alt. - destruct (ps !! j); [done|]. - iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } + iFrame. iRewrite "Heq" in "Hpj". + destruct (ps !! j); [done|]. by rewrite !option_equivI. } iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as (p2) "[Hm2 Hconsistent]". - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert; [done|]. lia. } + { rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert; [done|]. rewrite insert_length. lia. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame "Hm2". - iDestruct "Hi'" as %Hi. iDestruct "Hj'" as %Hj. + iDestruct "Hi'" as %Hi'. iDestruct "Hj'" as %Hj'. iSplitL "Hconsistent Hauth". { iExists (<[i:=p1]> (<[j:=p2]> ps)). iSplit. - { admit. - (* rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. *)} + { iPureIntro. rewrite !insert_length. done. } iFrame. rewrite list_insert_insert. rewrite list_insert_commute; [|done]. rewrite list_insert_insert. by rewrite list_insert_commute; [|done]. } iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. - Admitted. + Qed. Lemma iProto_target γ ps_dom i a j m : iProto_ctx γ ps_dom -∗ @@ -1265,9 +1343,9 @@ Section proto. iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } iSplitL "HSome". { iNext. iDestruct "HSome" as %Heq. - iPureIntro. simplify_eq. admit. } + iPureIntro. simplify_eq. by apply lookup_lt_is_Some_1. } iSplitL "Hauth Hps"; iExists _; by iFrame. - Admitted. + Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 9efa72bdd741c8aa00e88c4fc95c0bb39f38d361..e6b46d394fd1e033bcc423411bc64c780df3bb52 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -1,16 +1,15 @@ From multi_actris.channel Require Import proofmode. Set Default Proof Using "Type". -Definition iProto_empty {Σ} : gmap nat (iProto Σ) := ∅. +Definition iProto_empty {Σ} : list (iProto Σ) := []. Lemma iProto_consistent_empty {Σ} : ⊢ iProto_consistent (@iProto_empty Σ). -Proof. iProto_consistent_take_step. Qed. +Proof. iProto_consistent_take_steps. Qed. -Definition iProto_binary `{!heapGS Σ} : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto ]> - ∅). +Definition iProto_binary `{!heapGS Σ} : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto]. Lemma iProto_binary_consistent `{!heapGS Σ} : ⊢ iProto_consistent iProto_binary. @@ -29,10 +28,10 @@ Definition roundtrip_prog : val := Section channel. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_roundtrip : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). + Definition iProto_roundtrip : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto]. Lemma iProto_roundtrip_consistent : ⊢ iProto_consistent iProto_roundtrip. @@ -43,15 +42,12 @@ Section channel. {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip); - [lia|set_solver|iApply iProto_roundtrip_consistent|]. + wp_smart_apply (new_chan_spec iProto_roundtrip); + [set_solver|iApply iProto_roundtrip_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). @@ -74,13 +70,13 @@ Definition roundtrip_ref_prog : val := Section roundtrip_ref. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_roundtrip_ref : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> - (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). + Definition iProto_roundtrip_ref : list (iProto Σ) := + [(<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto; + (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto; + (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]. Lemma iProto_roundtrip_ref_consistent : ⊢ iProto_consistent iProto_roundtrip_ref. @@ -95,17 +91,12 @@ Section roundtrip_ref. {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref with "[]"). - { lia. } - { set_solver. } - { iApply iProto_roundtrip_ref_consistent. } + wp_smart_apply (new_chan_spec iProto_roundtrip_ref); + [set_solver|iApply iProto_roundtrip_ref_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } @@ -188,10 +179,10 @@ Section roundtrip_ref_rec. (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. - Definition iProto_roundtrip_ref_rec : gmap nat (iProto Σ) := - <[0 := iProto_roundtrip_ref_rec1]> - (<[1 := iProto_roundtrip_ref_rec2]> - (<[2 := iProto_roundtrip_ref_rec3]> ∅)). + Definition iProto_roundtrip_ref_rec : list (iProto Σ) := + [iProto_roundtrip_ref_rec1; + iProto_roundtrip_ref_rec2; + iProto_roundtrip_ref_rec3]. Lemma iProto_roundtrip_ref_rec_consistent : ⊢ iProto_consistent iProto_roundtrip_ref_rec. @@ -210,7 +201,6 @@ Section roundtrip_ref_rec. iIntros "Hloc". iSplit; [done|]. replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. rewrite -iProto_roundtrip_ref_rec2_unfold. - do 2 clean_map 0. do 2 clean_map 1. do 2 clean_map 2. iApply "IH". Qed. @@ -218,17 +208,12 @@ Section roundtrip_ref_rec. {{{ True }}} roundtrip_ref_rec_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref_rec with "[]"). - { lia. } - { set_solver. } - { iApply iProto_roundtrip_ref_rec_consistent. } + wp_smart_apply (new_chan_spec iProto_roundtrip_ref_rec); + [set_solver|iApply iProto_roundtrip_ref_rec_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_pure _. iLöb as "IH". wp_recv (l x) as "Hl". wp_load. wp_store. wp_send with "[$Hl]". @@ -258,13 +243,13 @@ Definition smuggle_ref_prog : val := Section smuggle_ref. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_smuggle_ref : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Recv,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; - <(Recv, 2)> MSG #(); <(Send,0)> MSG #() ; END)%proto]> - (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Send,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> ∅)). + Definition iProto_smuggle_ref : list (iProto Σ) := + [(<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Recv,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto; + (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; + <(Recv, 2)> MSG #(); <(Send,0)> MSG #() ; END)%proto; + (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Send,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]. Lemma iProto_smuggle_ref_consistent : ⊢ iProto_consistent iProto_smuggle_ref. @@ -274,18 +259,13 @@ Section smuggle_ref. {{{ True }}} smuggle_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_smuggle_ref with "[]"). - { lia. } - { set_solver. } - { iApply iProto_smuggle_ref_consistent. } + wp_smart_apply (new_chan_spec iProto_smuggle_ref); + [set_solver|iApply iProto_smuggle_ref_consistent|]. iIntros (cs) "Hcs". wp_pures. - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (v) as "_". wp_send with "[//]". wp_recv as "_". by wp_send with "[//]". } @@ -311,149 +291,57 @@ Section parallel. 0 *) - Definition iProto_parallel : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x1:Z)> MSG #x1 ; + Definition iProto_parallel : list (iProto Σ) := + [(<(Send, 1) @ (x1:Z)> MSG #x1 ; <(Send, 2) @ (x2:Z)> MSG #x2 ; <(Recv, 3) @ (y1:Z)> MSG #(x1+y1); - <(Recv, 4) @ (y2:Z)> MSG #(x2+y2); END)%proto]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 3) @ (y:Z)> MSG #(x+y); END)%proto]> - (<[2 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 4) @ (y:Z)> MSG #(x+y) ; END)%proto]> - (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x; END)%proto]> - (<[4 := (<(Recv, 2) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> ∅)))). + <(Recv, 4) @ (y2:Z)> MSG #(x2+y2); END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 3) @ (y:Z)> MSG #(x+y); END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 4) @ (y:Z)> MSG #(x+y) ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x; END)%proto; + (<(Recv, 2) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]. Lemma iProto_parallel_consistent : ⊢ iProto_consistent iProto_parallel. - Proof. - rewrite /iProto_parallel. - iProto_consistent_take_step. - iIntros (x1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - - iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - * iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 3. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - - iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - + iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros (z) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - Qed. + Proof. rewrite /iProto_parallel. iProto_consistent_take_steps. Qed. End parallel. Section two_buyer. Context `{!heapGS Σ}. - Definition two_buyer_prot : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (title:Z)> MSG #title ; + Definition two_buyer_prot : list (iProto Σ) := + [(<(Send, 1) @ (title:Z)> MSG #title ; <(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto]> - (<[1 := (<(Recv, 0) @ (title:Z)> MSG #title ; + <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto; + (<(Recv, 0) @ (title:Z)> MSG #title ; <(Send, 0) @ (quote:Z)> MSG #quote ; <(Send, 2)> MSG #quote ; <(Recv, 2) @ (b:bool)> MSG #b ; if b then <(Recv, 2) @ (address:Z)> MSG #address ; <(Send, 2) @ (date:Z)> MSG #date ; END - else END)%proto]> - (<[2 := (<(Recv, 1) @ (quote:Z)> MSG #quote ; + else END)%proto; + (<(Recv, 1) @ (quote:Z)> MSG #quote ; <(Recv, 0) @ (contrib:Z)> MSG #contrib ; if bool_decide (contrib >= quote/2)%Z then <(Send, 1)> MSG #true ; <(Send, 1) @ (address:Z)> MSG #address ; <(Recv, 1) @ (date:Z)> MSG #date ; END else - <(Send, 1)> MSG #false ; END)%proto]> - ∅)). + <(Send, 1)> MSG #false ; END)%proto]. Lemma two_buyer_prot_consistent : ⊢ iProto_consistent two_buyer_prot. Proof. - rewrite /two_buyer_prot. - iProto_consistent_take_step. - iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - iProto_consistent_take_step. - iIntros (contrib) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - case_bool_decide. - - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (address) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (date) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + case_bool_decide; iProto_consistent_take_steps. Qed. - + End two_buyer. Section two_buyer_ref. @@ -493,47 +381,21 @@ Section two_buyer_ref. END else END)%proto. - Definition two_buyer_ref_prot : gmap nat (iProto Σ) := - <[0 := two_buyer_ref_b1_prot]> - (<[1 := two_buyer_ref_s_prot]> - (<[2 := two_buyer_ref_b2_prot]> - ∅)). + Definition two_buyer_ref_prot : list (iProto Σ) := + [two_buyer_ref_b1_prot; + two_buyer_ref_s_prot; + two_buyer_ref_b2_prot]. + (* TODO: Anonymous variable in this is unsatisfactory *) Lemma two_buyer_ref_prot_consistent : ⊢ iProto_consistent two_buyer_ref_prot. Proof. - rewrite /two_buyer_prot. - iProto_consistent_take_step. - iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - iProto_consistent_take_step. - iIntros (l1 amount1 contrib) "Hl1". iExists _,_,_. iSplit; [done|]. iFrame. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - iIntros (b) "Hl1". iExists _. iSplit; [done|]. iFrame. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - destruct b. - - iProto_consistent_take_step. - iIntros (l2 amount2 address) "Hl2". iExists _,_,_. iSplit; [done|]. iFrame. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (date) "Hl2". iExists _. iSplit; [done|]. iFrame. - iProto_consistent_take_step. - - iProto_consistent_take_step. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + iNext. destruct x4; iProto_consistent_take_steps. Qed. End two_buyer_ref. - Section forwarder. Context `{!heapGS Σ}. @@ -548,36 +410,25 @@ Section forwarder. *) - Definition iProto_forwarder : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; + Definition iProto_forwarder : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 1) @ (b:bool)> MSG #b ; - <(Recv, if b then 2 else 3) > MSG #x ; END)%proto]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Recv, if b then 2 else 3) > MSG #x ; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; <(Recv, 0) @ (b:bool)> MSG #b; - <(Send, if b then 2 else 3)> MSG #x ; END)%proto]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> - (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> ∅))). + <(Send, if b then 2 else 3)> MSG #x ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]. + (* TODO: Anonymous variable in this is unsatisfactory *) Lemma iProto_forwarder_consistent : ⊢ iProto_consistent iProto_forwarder. Proof. rewrite /iProto_forwarder. - iProto_consistent_take_step. - iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros ([]) "_". - - iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 0. repeat clean_map 1. - iProto_consistent_take_steps. - - iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 0. repeat clean_map 1. - iProto_consistent_take_steps. + iProto_consistent_take_steps. + destruct x0; iProto_consistent_take_steps. Qed. End forwarder. @@ -637,11 +488,11 @@ Section forwarder_rec. (iProto_forwarder_rec_n_aux iProto_forwarder_rec_n). Proof. apply (fixpoint_unfold _). Qed. - Definition iProto_forwarder_rec : gmap nat (iProto Σ) := - <[0 := iProto_forwarder_rec_0]> - (<[1 := iProto_forwarder_rec_1]> - (<[2 := iProto_forwarder_rec_n]> - (<[3 := iProto_forwarder_rec_n]>∅))). + Definition iProto_forwarder_rec : list (iProto Σ) := + [iProto_forwarder_rec_0; + iProto_forwarder_rec_1; + iProto_forwarder_rec_n; + iProto_forwarder_rec_n]. Lemma iProto_forwarder_rec_consistent : ⊢ iProto_consistent iProto_forwarder_rec. @@ -665,7 +516,6 @@ Section forwarder_rec. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. - rewrite (insert_commute _ 2 1); [|done]. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). @@ -680,7 +530,6 @@ Section forwarder_rec. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. repeat clean_map 1. repeat clean_map 3. repeat clean_map 0. - rewrite (insert_commute _ 3 1); [|done]. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index cb21e823d4d481ec63a2fefc575d85f1000644e5..0a3bcff51cd0b837a3edc9bc4bd382250147a9c1 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -133,98 +133,54 @@ Section ring_leader_election_example. Definition prot_tail (i_max : nat) : iProto Σ := (<(Send,0)> MSG #i_max; END)%proto. - Definition pre_prot_pool id_max : gmap nat (iProto Σ) := - <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + Definition pre_prot_pool id_max : list (iProto Σ) := + [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; <(Recv,2)> MSG #id_max ; <(Recv,3)> MSG #id_max ; - END)%proto ]> - (<[1 := prot_tail id_max ]> - (<[2 := prot_tail id_max ]> - (<[3 := prot_tail id_max ]> ∅))). + END)%proto; + prot_tail id_max; + prot_tail id_max; + prot_tail id_max]. Lemma pre_prot_pool_consistent id_max : ⊢ iProto_consistent (pre_prot_pool id_max). Proof. rewrite /pre_prot_pool. iProto_consistent_take_steps. Qed. - Definition prot_pool : gmap nat (iProto Σ) := - <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + Definition prot_pool : list (iProto Σ) := + [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; <(Recv,2)> MSG #id_max ; <(Recv,3)> MSG #id_max ; - END)%proto ]> - (<[1 := rle_preprot 3 1 2 prot_tail ]> - (<[2 := rle_prot 1 2 3 prot_tail false ]> - (<[3 := rle_preprot 2 3 1 prot_tail ]> ∅))). + END)%proto; + rle_preprot 3 1 2 prot_tail; + rle_prot 1 2 3 prot_tail false; + rle_preprot 2 3 1 prot_tail]. Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. Proof. rewrite /prot_pool /rle_preprot. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_steps. + case_bool_decide; try lia. + case_bool_decide; try lia. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_steps. + case_bool_decide; try lia. + case_bool_decide; try lia. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - repeat (rewrite (insert_commute _ _ 3); [|lia]). - repeat (rewrite (insert_commute _ _ 2); [|lia]). - repeat (rewrite (insert_commute _ _ 1); [|lia]). - repeat (rewrite (insert_commute _ _ 0); [|lia]). - iApply pre_prot_pool_consistent. + iProto_consistent_take_steps. Qed. Lemma program_spec : {{{ True }}} program #() {{{ RET #(); True }}}. Proof. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 4 prot_pool); - [lia|set_solver|iApply prot_pool_consistent|]. + wp_smart_apply (new_chan_spec prot_pool); + [set_solver|iApply prot_pool_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [done|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [done|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [done|]. - iIntros (c2) "[Hc2 Hcs]". - wp_smart_apply (get_chan_spec _ 3 with "Hcs"); [done|]. - iIntros (c3) "[Hc3 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". + wp_get_chan (c3) "[Hc3 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). iIntros (i') "Hc1". by wp_send with "[//]". }