diff --git a/.gitlab-ci.yml b/.gitlab-ci.yml index 1b6b3e0d3b409687cdb0a26342febeba36761a20..1f03e57ad223b807c62d3309697b31d356134495 100644 --- a/.gitlab-ci.yml +++ b/.gitlab-ci.yml @@ -23,6 +23,7 @@ variables: except: - triggers - schedules + - api ## Build jobs @@ -43,3 +44,4 @@ build-iris.dev: only: - triggers - schedules + - api diff --git a/README.md b/README.md index 38bcaebb3c76a1a6aaffc957c3962281116e793f..2e942996acddfd03e322919032a52fdda6a19f81 100644 --- a/README.md +++ b/README.md @@ -53,8 +53,11 @@ This repository contains the following case studies: * [logrel_heaplang](theories/logrel_heaplang): A unary logical relation for semantic typing of heap lang. * [logatom](theories/logrel_heaplang): Proofs of various logically atomic specifications: - - Elimination Stack - - Treiber Stack (by Zhen Zhang) + - Elimination Stack (by Ralf Jung) + - Conditional increment (inspired by [this paper](https://people.mpi-sws.org/~dreyer/papers/relcon/paper.pdf)) and RDCSS (as in [this paper](https://timharris.uk/papers/2002-disc.pdf)) (by Marianna Rapoport, Rodolphe Lepigre and Gaurav Parthasarathy) + - [Herlihy-Wing-Queue](https://cs.brown.edu/~mph/HerlihyW90/p463-herlihy.pdf) + - Atomic Snapshot (by Marianna Rapoport) + - Treiber Stack (by Zhen Zhang, and another version by Rodolphe Lepigre) - Flat Combiner (by Zhen Zhang, also see [this archived documentation](https://gitlab.mpi-sws.org/FP/iris-atomic/tree/master/docs)) * [spanning-tree](theories/spanning_tree): Proof of a concurrent spanning tree algorithm by Amin Timany. diff --git a/_CoqProject b/_CoqProject index 3d6340ccc500d1fc45f5c580e9ce4a3a9ca3f367..467617a6d3947e8c8542208e673aa3b8e6a95a1c 100644 --- a/_CoqProject +++ b/_CoqProject @@ -93,6 +93,7 @@ theories/hocap/lib/oneshot.v theories/hocap/concurrent_runners.v theories/hocap/parfib.v +theories/logatom/lib/gc.v theories/logatom/treiber.v theories/logatom/treiber2.v theories/logatom/elimination_stack/hocap_spec.v @@ -108,3 +109,8 @@ theories/logatom/snapshot/spec.v theories/logatom/snapshot/atomic_snapshot.v theories/logatom/conditional_increment/spec.v theories/logatom/conditional_increment/cinc.v +theories/logatom/rdcss/rdcss.v +theories/logatom/rdcss/spec.v +theories/logatom/proph_erasure.v +theories/logatom/herlihy_wing_queue/spec.v +theories/logatom/herlihy_wing_queue/hwq.v diff --git a/opam b/opam index ec4e7bc8ee89ee4a77c11ecb371ff376b9c473be..b481c5018a61d6f308d28d48609ff38733d08ae3 100644 --- a/opam +++ b/opam @@ -1,7 +1,7 @@ opam-version: "1.2" name: "coq-iris-examples" maintainer: "Ralf Jung " -authors: "The Iris Team" +authors: "The Iris Team and Contributors" homepage: "http://iris-project.org/" bug-reports: "https://gitlab.mpi-sws.org/FP/iris-examples/issues" dev-repo: "https://gitlab.mpi-sws.org/FP/iris-examples.git" @@ -9,6 +9,6 @@ build: [make "-j%{jobs}%"] install: [make "install"] remove: ["rm" "-rf" "%{lib}%/coq/user-contrib/iris_examples"] depends: [ - "coq-iris" { (= "dev.2019-06-11.8.a51fa3cf") | (= "dev") } + "coq-iris" { (= "dev.2019-08-14.0.ffccb508") | (= "dev") } "coq-autosubst" { = "dev.coq86" } ] diff --git a/theories/barrier/example_joining_existentials.v b/theories/barrier/example_joining_existentials.v index 40cd79b36b6d4486dfaa5e9ebf2410701e87e814..0ed7746c94d6abf3b6c5f66fe497f864ad0490e1 100644 --- a/theories/barrier/example_joining_existentials.v +++ b/theories/barrier/example_joining_existentials.v @@ -6,16 +6,16 @@ From iris.proofmode Require Import tactics. From iris_examples.barrier Require Import proof specification. Set Default Proof Using "Type". -Definition one_shotR (Σ : gFunctors) (F : cFunctor) := - csumR (exclR unitC) (agreeR $ laterC $ F (iPreProp Σ) _). +Definition one_shotR (Σ : gFunctors) (F : oFunctor) := + csumR (exclR unitO) (agreeR $ laterO $ F (iPreProp Σ) _). Definition Pending {Σ F} : one_shotR Σ F := Cinl (Excl ()). -Definition Shot {Σ} {F : cFunctor} (x : F (iProp Σ) _) : one_shotR Σ F := - Cinr $ to_agree $ Next $ cFunctor_map F (iProp_fold, iProp_unfold) x. +Definition Shot {Σ} {F : oFunctor} (x : F (iProp Σ) _) : one_shotR Σ F := + Cinr $ to_agree $ Next $ oFunctor_map F (iProp_fold, iProp_unfold) x. -Class oneShotG (Σ : gFunctors) (F : cFunctor) := +Class oneShotG (Σ : gFunctors) (F : oFunctor) := one_shot_inG :> inG Σ (one_shotR Σ F). -Definition oneShotΣ (F : cFunctor) : gFunctors := - #[ GFunctor (csumRF (exclRF unitC) (agreeRF (▶ F))) ]. +Definition oneShotΣ (F : oFunctor) : gFunctors := + #[ GFunctor (csumRF (exclRF unitO) (agreeRF (▶ F))) ]. Instance subG_oneShotΣ {Σ F} : subG (oneShotΣ F) Σ → oneShotG Σ F. Proof. solve_inG. Qed. @@ -59,12 +59,12 @@ Proof. iAssert (▷ (x ≡ x'))%I as "Hxx". { iCombine "Hγ" "Hγ'" as "Hγ2". iClear "Hγ Hγ'". rewrite own_valid csum_validI /= agree_validI agree_equivI bi.later_equivI /=. - rewrite -{2}[x]cFunctor_id -{2}[x']cFunctor_id. - assert (HF : cFunctor_map F (cid, cid) ≡ cFunctor_map F (iProp_fold (Σ:=Σ) ◎ iProp_unfold, iProp_fold (Σ:=Σ) ◎ iProp_unfold)). + rewrite -{2}[x]oFunctor_id -{2}[x']oFunctor_id. + assert (HF : oFunctor_map F (cid, cid) ≡ oFunctor_map F (iProp_fold (Σ:=Σ) ◎ iProp_unfold, iProp_fold (Σ:=Σ) ◎ iProp_unfold)). { apply ne_proper; first by apply _. by split; intro; simpl; symmetry; apply iProp_fold_unfold. } rewrite (HF x). rewrite (HF x'). - rewrite !cFunctor_compose. iNext. by iRewrite "Hγ2". } + rewrite !oFunctor_compose. iNext. by iRewrite "Hγ2". } iNext. iRewrite -"Hxx" in "Hx'". iExists x; iFrame "Hγ". iApply (Ψ_join with "Hx Hx'"). Qed. diff --git a/theories/barrier/specification.v b/theories/barrier/specification.v index 0dceb3e6f1f761e228d79d32686041e8b6274969..650e001176a3b5e1d3af6ca11454903e78dcdad8 100644 --- a/theories/barrier/specification.v +++ b/theories/barrier/specification.v @@ -18,7 +18,7 @@ Lemma barrier_spec (N : namespace) : (∀ l P Q, recv l (P ∗ Q) ={↑N}=> recv l P ∗ recv l Q) ∧ (∀ l P Q, (P -∗ Q) -∗ recv l P -∗ recv l Q). Proof. - exists (λ l, CofeMor (recv N l)), (λ l, CofeMor (send N l)). + exists (λ l, OfeMor (recv N l)), (λ l, OfeMor (send N l)). split_and?; simpl. - iIntros (P) "!# _". iApply (newbarrier_spec _ P with "[]"); [done..|]. iNext. eauto. diff --git a/theories/concurrent_stacks/concurrent_stack1.v b/theories/concurrent_stacks/concurrent_stack1.v index f471e164f8a1269d69af9db0f8d36c227f31abc2..7fbdcc241f3fa4fd7318586da75e60d616a4533b 100644 --- a/theories/concurrent_stacks/concurrent_stack1.v +++ b/theories/concurrent_stacks/concurrent_stack1.v @@ -36,9 +36,19 @@ Section stacks. iIntros "H"; iDestruct "H" as (?) "[Hl Hl']"; iSplitL "Hl"; eauto. Qed. - Definition is_list_pre (P : val → iProp Σ) (F : val -c> iProp Σ) : - val -c> iProp Σ := λ v, - (v ≡ NONEV ∨ ∃ (l : loc) (h t : val), ⌜v ≡ SOMEV #l⌝ ∗ l ↦{-} (h, t)%V ∗ P h ∗ ▷ F t)%I. + Definition oloc_to_val (ol: option loc) : val := + match ol with + | None => NONEV + | Some loc => SOMEV (#loc) + end. + Local Instance oloc_to_val_inj : Inj (=) (=) oloc_to_val. + Proof. intros [|][|]; simpl; congruence. Qed. + + Definition is_list_pre (P : val → iProp Σ) (F : option loc -d> iProp Σ) : + option loc -d> iProp Σ := λ v, match v with + | None => True + | Some l => ∃ (h : val) (t : option loc), l ↦{-} (h, oloc_to_val t)%V ∗ P h ∗ ▷ F t + end%I. Local Instance is_list_contr (P : val → iProp Σ) : Contractive (is_list_pre P). Proof. @@ -58,28 +68,22 @@ Section stacks. rewrite is_list_eq. apply (fixpoint_unfold (is_list_pre P)). Qed. - (* TODO: shouldn't have to explicitly return is_list *) - Lemma is_list_unboxed (P : val → iProp Σ) v : - is_list P v -∗ ⌜val_is_unboxed v⌝ ∗ is_list P v. - Proof. - iIntros "Hstack"; iSplit; last done; - iDestruct (is_list_unfold with "Hstack") as "[->|Hstack]"; - last iDestruct "Hstack" as (l h t) "(-> & _)"; done. - Qed. - - Lemma is_list_disj (P : val → iProp Σ) v : - is_list P v -∗ is_list P v ∗ (⌜v ≡ NONEV⌝ ∨ ∃ (l : loc) h t, ⌜v ≡ SOMEV #l%V⌝ ∗ l ↦{-} (h, t)%V). + Lemma is_list_dup (P : val → iProp Σ) v : + is_list P v -∗ is_list P v ∗ match v with + | None => True + | Some l => ∃ h t, l ↦{-} (h, oloc_to_val t)%V + end. Proof. - iIntros "Hstack". - iDestruct (is_list_unfold with "Hstack") as "[%|Hstack]"; simplify_eq. - - rewrite is_list_unfold; iSplitR; [iLeft|]; eauto. - - iDestruct "Hstack" as (l h t) "(% & Hl & Hlist)". - iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]"; simplify_eq. - rewrite (is_list_unfold _ (InjRV _)); iSplitR "Hl2"; iRight; iExists _, _, _; by iFrame. + iIntros "Hstack". iDestruct (is_list_unfold with "Hstack") as "Hstack". + destruct v as [l|]. + - iDestruct "Hstack" as (h t) "(Hl & Hlist)". + iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]". + rewrite (is_list_unfold _ (Some _)); iSplitR "Hl2"; iExists _, _; by iFrame. + - rewrite is_list_unfold; iSplitR; eauto. Qed. Definition stack_inv P v := - (∃ l v', ⌜v = #l⌝ ∗ l ↦ v' ∗ is_list P v')%I. + (∃ l ol', ⌜v = #l⌝ ∗ l ↦ oloc_to_val ol' ∗ is_list P ol')%I. Definition is_stack (P : val → iProp Σ) v := inv N (stack_inv P v). @@ -92,8 +96,8 @@ Section stacks. wp_lam. wp_alloc ℓ as "Hl". iMod (inv_alloc N ⊤ (stack_inv P #ℓ) with "[Hl]") as "Hinv". - { iNext; iExists ℓ, NONEV; iFrame; - by iSplit; last (iApply is_list_unfold; iLeft). } + { iNext; iExists ℓ, None; iFrame; + by iSplit; last (iApply is_list_unfold). } by iApply "Hpost". Qed. @@ -107,23 +111,23 @@ Section stacks. wp_load. iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _, _; by iFrame. } - iModIntro. wp_let. wp_alloc ℓ' as "Hl'". wp_pures. wp_bind (CAS _ _ _). + iModIntro. wp_let. wp_alloc ℓ' as "Hl'". wp_pures. wp_bind (CmpXchg _ _ _). iInv N as (ℓ'' v'') "(>% & >Hl & Hlist)" "Hclose"; simplify_eq. - destruct (decide (v' = v'')) as [ -> |]. - - iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". - wp_cas_suc. + destruct (decide (v' = v'')) as [->|Hne]. + - wp_cmpxchg_suc. { destruct v''; left; done. } iMod ("Hclose" with "[HP Hl Hl' Hlist]") as "_". - { iNext; iExists _, (InjRV #ℓ'); iFrame; iSplit; first done; - rewrite (is_list_unfold _ (InjRV _)). iRight; iExists _, _, _; iFrame; eauto. } + { iNext; iExists _, (Some ℓ'); iFrame; iSplit; first done; + rewrite (is_list_unfold _ (Some _)). iExists _, _; iFrame; eauto. } iModIntro. - wp_if. + wp_pures. by iApply "HΦ". - - iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". - wp_cas_fail. + - wp_cmpxchg_fail. + { destruct v', v''; simpl; congruence. } + { destruct v''; left; done. } iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _, _; by iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HP HΦ"). Qed. @@ -134,41 +138,40 @@ Section stacks. iLöb as "IH". wp_lam. wp_bind (Load _). iInv N as (ℓ v') "(>% & Hl & Hlist)" "Hclose"; subst. + iDestruct (is_list_dup with "Hlist") as "[Hlist Hlist2]". wp_load. - iDestruct (is_list_disj with "Hlist") as "[Hlist Hdisj]". iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _, _; by iFrame. } iModIntro. - iDestruct "Hdisj" as "[-> | Heq]". + destruct v' as [l|]; last first. - wp_match. iApply "HΦ"; by iLeft. - - iDestruct "Heq" as (l h t) "[-> Hl]". - wp_match. wp_bind (Load _). + - wp_match. wp_bind (Load _). iInv N as (ℓ' v') "(>% & Hl' & Hlist)" "Hclose". simplify_eq. - iDestruct "Hl" as (q) "Hl". + iDestruct "Hlist2" as (???) "Hl". wp_load. iMod ("Hclose" with "[Hl' Hlist]") as "_". { iNext; iExists _, _; by iFrame. } iModIntro. - wp_pures. wp_bind (CAS _ _ _). + wp_pures. wp_bind (CmpXchg _ _ _). iInv N as (ℓ'' v'') "(>% & Hl' & Hlist)" "Hclose". simplify_eq. - destruct (decide (v'' = InjRV #l)) as [-> |]. + destruct (decide (v'' = (Some l))) as [-> |]. * rewrite is_list_unfold. - iDestruct "Hlist" as "[>% | H]"; first done. - iDestruct "H" as (ℓ''' h' t') "(>% & Hl'' & HP & Hlist)"; simplify_eq. + iDestruct "Hlist" as (h' t') "(Hl'' & HP & Hlist)". iDestruct "Hl''" as (q') "Hl''". - wp_cas_suc. - iDestruct (mapsto_agree with "Hl'' Hl") as "%"; simplify_eq. + simpl. + wp_cmpxchg_suc. + iDestruct (mapsto_agree with "Hl'' Hl") as %[= <- <-%oloc_to_val_inj]. iMod ("Hclose" with "[Hl' Hlist]") as "_". { iNext; iExists ℓ'', _; by iFrame. } iModIntro. wp_pures. - iApply ("HΦ" with "[HP]"); iRight; iExists h; by iFrame. - * wp_cas_fail. + iApply ("HΦ" with "[HP]"); iRight; iExists _; by iFrame. + * wp_cmpxchg_fail. { destruct v''; simpl; congruence. } iMod ("Hclose" with "[Hl' Hlist]") as "_". { iNext; iExists ℓ'', _; by iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HΦ"). Qed. End stacks. diff --git a/theories/concurrent_stacks/concurrent_stack2.v b/theories/concurrent_stacks/concurrent_stack2.v index 7f11d03c2e9b672826d4f2918701f15a7d85265d..04703703bba5687d795774f2a767796ded21d8d5 100644 --- a/theories/concurrent_stacks/concurrent_stack2.v +++ b/theories/concurrent_stacks/concurrent_stack2.v @@ -97,22 +97,22 @@ Section side_channel. {{{ v', RET v'; (∃ v'' : val, ⌜v' = InjRV v''⌝ ∗ P v'') ∨ ⌜v' = InjLV #()⌝ }}}. Proof. iIntros (Φ) "[Hinv Hγ] HΦ". iDestruct "Hinv" as (v' l) "[-> #Hinv]". - wp_lam. wp_bind (CAS _ _ _). wp_pures. + wp_lam. wp_bind (CmpXchg _ _ _). wp_pures. iInv N as "Hstages" "Hclose". iDestruct "Hstages" as "[[Hl HP] | [H | [Hl H]]]". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hclose" with "[Hl Hγ]") as "_". { iRight; iRight; iFrame. } iModIntro. wp_pures. by iApply "HΦ"; iLeft; iExists _; iSplit. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[H]") as "_". { iRight; iLeft; auto. } iModIntro. wp_pures. by iApply "HΦ"; iRight. - - wp_cas_fail. + - wp_cmpxchg_fail. iDestruct (own_valid_2 with "H Hγ") as %[]. Qed. @@ -123,22 +123,22 @@ Section side_channel. {{{ v', RET v'; (∃ v'' : val, ⌜v' = InjRV v''⌝ ∗ P v'') ∨ ⌜v' = InjLV #()⌝ }}}. Proof. iIntros (Φ) "H HΦ"; iDestruct "H" as (v l) "[-> #Hinv]". - wp_lam. wp_proj. wp_bind (CAS _ _ _). + wp_lam. wp_proj. wp_bind (CmpXchg _ _ _). iInv N as "Hstages" "Hclose". iDestruct "Hstages" as "[[H HP] | [H | [Hl Hγ]]]". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hclose" with "[H]") as "_". { by iRight; iLeft. } iModIntro. wp_pures. iApply "HΦ"; iLeft; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[H]") as "_". { by iRight; iLeft. } iModIntro. wp_pures. iApply "HΦ"; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hl Hγ]"). { iRight; iRight; iFrame. } iModIntro. @@ -246,9 +246,19 @@ Section stack_works. iIntros "H"; iDestruct "H" as (?) "[Hl Hl']"; iSplitL "Hl"; eauto. Qed. - Definition is_list_pre (P : val → iProp Σ) (F : val -c> iProp Σ) : - val -c> iProp Σ := λ v, - (v ≡ NONEV ∨ ∃ (l : loc) (h t : val), ⌜v ≡ SOMEV #l⌝ ∗ l ↦{-} (h, t)%V ∗ P h ∗ ▷ F t)%I. + Definition oloc_to_val (ol: option loc) : val := + match ol with + | None => NONEV + | Some loc => SOMEV (#loc) + end. + Local Instance oloc_to_val_inj : Inj (=) (=) oloc_to_val. + Proof. intros [|][|]; simpl; congruence. Qed. + + Definition is_list_pre (P : val → iProp Σ) (F : option loc -d> iProp Σ) : + option loc -d> iProp Σ := λ v, match v with + | None => True + | Some l => ∃ (h : val) (t : option loc), l ↦{-} (h, oloc_to_val t)%V ∗ P h ∗ ▷ F t + end%I. Local Instance is_list_contr (P : val → iProp Σ) : Contractive (is_list_pre P). Proof. @@ -268,27 +278,21 @@ Section stack_works. rewrite is_list_eq. apply (fixpoint_unfold (is_list_pre P)). Qed. - (* TODO: shouldn't have to explicitly return is_list *) - Lemma is_list_unboxed (P : val → iProp Σ) v : - is_list P v -∗ ⌜val_is_unboxed v⌝ ∗ is_list P v. - Proof. - iIntros "Hstack"; iSplit; last done; - iDestruct (is_list_unfold with "Hstack") as "[->|Hstack]"; - last iDestruct "Hstack" as (l h t) "(-> & _)"; done. - Qed. - - Lemma is_list_disj (P : val → iProp Σ) v : - is_list P v -∗ is_list P v ∗ (⌜v ≡ NONEV⌝ ∨ ∃ (l : loc) h t, ⌜v ≡ SOMEV #l%V⌝ ∗ l ↦{-} (h, t)%V). + Lemma is_list_dup (P : val → iProp Σ) v : + is_list P v -∗ is_list P v ∗ match v with + | None => True + | Some l => ∃ h t, l ↦{-} (h, oloc_to_val t)%V + end. Proof. - iIntros "Hstack". - iDestruct (is_list_unfold with "Hstack") as "[%|Hstack]"; simplify_eq. - - rewrite is_list_unfold; iSplitR; [iLeft|]; eauto. - - iDestruct "Hstack" as (l h t) "(% & Hl & Hlist)". - iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]"; simplify_eq. - rewrite (is_list_unfold _ (InjRV _)); iSplitR "Hl2"; iRight; iExists _, _, _; by iFrame. + iIntros "Hstack". iDestruct (is_list_unfold with "Hstack") as "Hstack". + destruct v as [l|]. + - iDestruct "Hstack" as (h t) "(Hl & Hlist)". + iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]". + rewrite (is_list_unfold _ (Some _)); iSplitR "Hl2"; iExists _, _; by iFrame. + - rewrite is_list_unfold; iSplitR; eauto. Qed. - Definition stack_inv P l := (∃ v, l ↦ v ∗ is_list P v)%I. + Definition stack_inv P l := (∃ v, l ↦ oloc_to_val v ∗ is_list P v)%I. Definition is_stack P v := (∃ mailbox l, ⌜v = (mailbox, #l)%V⌝ ∗ is_mailbox Nmailbox P mailbox ∗ inv N (stack_inv P l))%I. @@ -302,7 +306,7 @@ Section stack_works. wp_apply mk_mailbox_works; first done. iIntros (mailbox) "#Hmailbox". iMod (inv_alloc N _ (stack_inv P l) with "[Hl]") as "#Hinv". - { by iNext; iExists _; iFrame; rewrite is_list_unfold; iLeft. } + { iNext; iExists None; iFrame. rewrite is_list_unfold. done. } wp_pures; iModIntro; iApply "Hpost"; iExists _, _; auto. Qed. @@ -322,23 +326,23 @@ Section stack_works. iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; iFrame. } iModIntro. - wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CAS _ _ _). + wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CmpXchg _ _ _). iInv N as (list) "(Hl & Hlist)" "Hclose". destruct (decide (v'' = list)) as [ -> |]. - * iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". - wp_cas_suc. + * wp_cmpxchg_suc. { destruct list; left; done. } iMod ("Hclose" with "[HP Hl Hl' Hlist]") as "_". - { iNext; iExists (SOMEV _); iFrame. - rewrite (is_list_unfold _ (InjRV _)). iRight; iExists _, _, _; iFrame; eauto. } + { iNext; iExists (Some _); iFrame. + rewrite (is_list_unfold _ (Some _)). iExists _, _; iFrame; eauto. } iModIntro. - wp_if. + wp_pures. by iApply "HΦ". - * iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". - wp_cas_fail. + * wp_cmpxchg_fail. + { destruct list, v''; simpl; congruence. } + { destruct list; left; done. } iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; by iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HP HΦ"). - wp_match. by iApply "HΦ". @@ -355,40 +359,38 @@ Section stack_works. - wp_match. wp_bind (Load _). iInv N as (list) "[Hl Hlist]" "Hclose". wp_load. - iDestruct (is_list_disj with "Hlist") as "[Hlist Hdisj]". + iDestruct (is_list_dup with "Hlist") as "[Hlist Hlist2]". iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; by iFrame. } iModIntro. - iDestruct "Hdisj" as "[-> | Heq]". + destruct list as [list|]; last first. * wp_match. iApply "HΦ"; by iLeft. - * iDestruct "Heq" as (l' h t) "[-> Hl']". - wp_match. wp_bind (Load _). + * wp_match. wp_bind (Load _). iInv N as (v') "[>Hl Hlist]" "Hclose". - iDestruct "Hl'" as (q) "Hl'". + iDestruct "Hlist2" as (???) "Hl'". wp_load. iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; by iFrame. } iModIntro. - wp_let. wp_proj. wp_bind (CAS _ _ _). wp_pures. + wp_let. wp_proj. wp_bind (CmpXchg _ _ _). wp_pures. iInv N as (v'') "[Hl Hlist]" "Hclose". - destruct (decide (v'' = InjRV #l')) as [-> |]. + destruct (decide (v'' = Some list)) as [-> |]. + rewrite is_list_unfold. - iDestruct "Hlist" as "[>% | H]"; first done. - iDestruct "H" as (l'' h' t') "(>% & Hl'' & HP & Hlist)"; simplify_eq. + iDestruct "Hlist" as (h' t') "(Hl'' & HP & Hlist)". iDestruct "Hl''" as (q') "Hl''". - wp_cas_suc. + wp_cmpxchg_suc. iDestruct (mapsto_agree with "Hl'' Hl'") as "%"; simplify_eq. iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; by iFrame. } iModIntro. wp_pures. - iApply ("HΦ" with "[HP]"); iRight; iExists h; by iFrame. - + wp_cas_fail. + iApply ("HΦ" with "[HP]"); iRight; iExists _; by iFrame. + + wp_cmpxchg_fail. { destruct v''; simpl; congruence. } iMod ("Hclose" with "[Hl Hlist]") as "_". { iNext; iExists _; by iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HΦ"). - iDestruct "HSome" as (v) "[-> HP]". wp_pures. diff --git a/theories/concurrent_stacks/concurrent_stack3.v b/theories/concurrent_stacks/concurrent_stack3.v index 33201c1d23edc8ace676830bf07ef708061bfdfc..5a3506bbbc4bd4b16edd7e993ffd8004d9bd6fd1 100644 --- a/theories/concurrent_stacks/concurrent_stack3.v +++ b/theories/concurrent_stacks/concurrent_stack3.v @@ -44,47 +44,51 @@ Section stack_works. iApply (mapsto_agree with "H1 H2"). Qed. + Definition oloc_to_val (ol: option loc) : val := + match ol with + | None => NONEV + | Some loc => SOMEV (#loc) + end. + Local Instance oloc_to_val_inj : Inj (=) (=) oloc_to_val. + Proof. intros [|][|]; simpl; congruence. Qed. + Fixpoint is_list xs v : iProp Σ := - (match xs with - | [] => ⌜v = NONEV⌝ - | x :: xs => ∃ l (t : val), ⌜v = SOMEV #l%V⌝ ∗ l ↦{-} (x, t)%V ∗ is_list xs t + (match xs, v with + | [], None => True + | x :: xs, Some l => ∃ t, l ↦{-} (x, oloc_to_val t)%V ∗ is_list xs t + | _, _ => False end)%I. - Lemma is_list_disj xs v : - is_list xs v -∗ is_list xs v ∗ (⌜v = NONEV⌝ ∨ ∃ l (h t : val), ⌜v = SOMEV #l⌝ ∗ l ↦{-} (h, t)%V). + Lemma is_list_dup xs v : + is_list xs v -∗ is_list xs v ∗ match v with + | None => True + | Some l => ∃ h t, l ↦{-} (h, oloc_to_val t)%V + end. Proof. - destruct xs; auto. - iIntros "H"; iDestruct "H" as (l t) "(-> & Hl & Hstack)". + destruct xs, v; simpl; auto; first by iIntros "[]". + iIntros "H"; iDestruct "H" as (t) "(Hl & Hstack)". iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]". - iSplitR "Hl2"; first by (iExists _, _; iFrame). iRight; auto. - Qed. - - Lemma is_list_unboxed xs v : - is_list xs v -∗ ⌜val_is_unboxed v⌝ ∗ is_list xs v. - Proof. - iIntros "Hlist"; iDestruct (is_list_disj with "Hlist") as "[$ Heq]". - iDestruct "Heq" as "[-> | H]"; first done; by iDestruct "H" as (? ? ?) "[-> ?]". + iSplitR "Hl2"; first by (iExists _; iFrame). by iExists _, _. Qed. Lemma is_list_empty xs : - is_list xs (InjLV #()) -∗ ⌜xs = []⌝. + is_list xs None -∗ ⌜xs = []⌝. Proof. destruct xs; iIntros "Hstack"; auto. - iDestruct "Hstack" as (? ?) "(% & H)"; discriminate. Qed. Lemma is_list_cons xs l h t : l ↦{-} (h, t)%V -∗ - is_list xs (InjRV #l) -∗ + is_list xs (Some l) -∗ ∃ ys, ⌜xs = h :: ys⌝. Proof. destruct xs; first by iIntros "? %". - iIntros "Hl Hstack"; iDestruct "Hstack" as (l' t') "(% & Hl' & Hrest)"; simplify_eq. + iIntros "Hl Hstack"; iDestruct "Hstack" as (t') "(Hl' & Hrest)". iDestruct (partial_mapsto_agree with "Hl Hl'") as "%"; simplify_eq; iExists _; auto. Qed. Definition stack_inv P l := - (∃ v xs, l ↦ v ∗ is_list xs v ∗ P xs)%I. + (∃ v xs, l ↦ oloc_to_val v ∗ is_list xs v ∗ P xs)%I. Definition is_stack_pred P v := (∃ l, ⌜v = #l⌝ ∗ inv N (stack_inv P l))%I. @@ -96,7 +100,7 @@ Section stack_works. rewrite -wp_fupd. wp_lam. wp_alloc l as "Hl". iMod (inv_alloc N _ (stack_inv P l) with "[Hl HP]") as "#Hinv". - { by iNext; iExists _, []; iFrame. } + { iNext; iExists None, []; iFrame. } iModIntro; iApply "HΦ"; iExists _; auto. Qed. @@ -114,22 +118,23 @@ Section stack_works. { iNext; iExists _, _; iFrame. } clear xs. iModIntro. - wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CAS _ _ _). + wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CmpXchg _ _ _). iInv N as (list' xs) "(Hl & Hlist & HP)" "Hclose". - iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". destruct (decide (list = list')) as [ -> |]. - - wp_cas_suc. + - wp_cmpxchg_suc. { destruct list'; left; done. } iMod ("Hupd" with "HP") as "[HP HΨ]". iMod ("Hclose" with "[Hl Hl' HP Hlist]") as "_". - { iNext; iExists (SOMEV _), (v :: xs); iFrame; iExists _, _; iFrame; auto. } + { iNext; iExists (Some _), (v :: xs); iFrame; iExists _; iFrame; auto. } iModIntro. - wp_if. + wp_pures. by iApply ("HΦ" with "HΨ"). - - wp_cas_fail. + - wp_cmpxchg_fail. + { destruct list, list'; simpl; congruence. } + { destruct list'; left; done. } iMod ("Hclose" with "[Hl HP Hlist]"). { iExists _, _; iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "Hupd HΦ"). Qed. @@ -146,8 +151,8 @@ Section stack_works. wp_lam. wp_bind (Load _). iInv N as (v xs) "(Hl & Hlist & HP)" "Hclose". wp_load. - iDestruct (is_list_disj with "Hlist") as "[Hlist H]". - iDestruct "H" as "[-> | HSome]". + iDestruct (is_list_dup with "Hlist") as "[Hlist H]". + destruct v as [l'|]; last first. - iDestruct (is_list_empty with "Hlist") as %->. iDestruct "Hupd" as "[_ Hupdnil]". iMod ("Hupdnil" with "HP") as "[HP HΨ]". @@ -156,7 +161,7 @@ Section stack_works. iModIntro. wp_match. iApply ("HΦ" with "HΨ"). - - iDestruct "HSome" as (l' h t) "[-> Hl']". + - iDestruct "H" as (h t) "Hl'". iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. @@ -167,15 +172,15 @@ Section stack_works. iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. - wp_let. wp_proj. wp_bind (CAS _ _ _). wp_pures. + wp_let. wp_proj. wp_bind (CmpXchg _ _ _). wp_pures. iInv N as (v' xs'') "(Hl & Hlist & HP)" "Hclose". - destruct (decide (v' = (SOMEV #l'))) as [ -> |]. - * wp_cas_suc. + destruct (decide (v' = (Some l'))) as [ -> |]. + * wp_cmpxchg_suc. iDestruct (is_list_cons with "[Hl'] Hlist") as (ys) "%"; first by iExists _. simplify_eq. iDestruct "Hupd" as "[Hupdcons _]". iMod ("Hupdcons" with "HP") as "[HP HΨ]". - iDestruct "Hlist" as (l'' t') "(% & Hl'' & Hlist)"; simplify_eq. + iDestruct "Hlist" as (t') "(Hl'' & Hlist)". iDestruct "Hl''" as (q') "Hl''". iDestruct (mapsto_agree with "Hl' Hl''") as "%"; simplify_eq. iMod ("Hclose" with "[Hlist Hl HP]") as "_". @@ -183,11 +188,11 @@ Section stack_works. iModIntro. wp_pures. iApply ("HΦ" with "HΨ"). - * wp_cas_fail. + * wp_cmpxchg_fail. { destruct v'; simpl; congruence. } iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "Hupd HΦ"). Qed. End stack_works. diff --git a/theories/concurrent_stacks/concurrent_stack4.v b/theories/concurrent_stacks/concurrent_stack4.v index 3265ef23658457c5c55bef33d6b7f48606a9f83d..57af553ed617bc80883cd4bc41fd1bb1faeb3094 100644 --- a/theories/concurrent_stacks/concurrent_stack4.v +++ b/theories/concurrent_stacks/concurrent_stack4.v @@ -119,24 +119,24 @@ Section proofs. {{{ v', RET v'; (∃ v'' : val, ⌜v' = InjRV v''⌝ ∗ can_push P Q v'') ∨ (⌜v' = InjLV #()⌝ ∗ (Q #())) }}}. Proof. iIntros (Φ) "[Hinv Hγ] HΦ". iDestruct "Hinv" as (v' l) "[-> #Hinv]". - wp_lam. wp_pures. wp_bind (CAS _ _ _). + wp_lam. wp_pures. wp_bind (CmpXchg _ _ _). iInv Nside_channel as "Hstages" "Hclose". iDestruct "Hstages" as "[[Hl HP] | [[Hl HQ] | [[Hl H] | [Hl H]]]]". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hclose" with "[Hl Hγ]") as "_". { iNext; iRight; iRight; iFrame. } iModIntro. wp_pures. by iApply "HΦ"; iLeft; iExists _; iFrame. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hl Hγ]") as "_". { iNext; iRight; iRight; iLeft; iFrame. } iModIntro. wp_pures. iApply ("HΦ" with "[HQ]"); iRight; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iDestruct (own_valid_2 with "H Hγ") as %[]. - - wp_cas_fail. + - wp_cmpxchg_fail. iDestruct (own_valid_2 with "H Hγ") as %[]. Qed. @@ -149,11 +149,11 @@ Section proofs. (∃ v'' : val, ⌜v' = InjRV v''⌝ ∗ Ψ v') ∨ (⌜v' = InjLV #()⌝ ∗ (do_pop ∧ Q')) }}}. Proof. simpl; iIntros (Φ) "[H [Hopener Hupd]] HΦ"; iDestruct "H" as (v l) "[-> #Hinv]". - wp_lam. wp_proj. wp_bind (CAS _ _ _). + wp_lam. wp_proj. wp_bind (CmpXchg _ _ _). iInv Nside_channel as "Hstages" "Hclose". iDestruct "Hstages" as "[[Hl Hpush] | [[Hl HQ] | [[Hl Hγ] | [Hl Hγ]]]]". - iMod "Hopener" as (xs) "[HP Hcloser]". - wp_cas_suc. + wp_cmpxchg_suc. iMod ("Hpush" with "HP") as "[HP HQ]". iMod ("Hupd" with "HP") as "[HP HΨ]". iMod ("Hcloser" with "HP") as "_". @@ -162,19 +162,19 @@ Section proofs. iApply fupd_intro_mask; first done. wp_pures. iApply "HΦ"; iLeft; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hl HQ]") as "_". { iRight; iLeft; iFrame. } iModIntro. wp_pures. iApply "HΦ"; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hl Hγ]"). { iRight; iRight; iFrame. } iModIntro. wp_pures. iApply "HΦ"; auto. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hl Hγ]"). { iRight; iRight; iFrame. } iModIntro. @@ -267,47 +267,51 @@ Section proofs. iApply (mapsto_agree with "H1 H2"). Qed. + Definition oloc_to_val (ol: option loc) : val := + match ol with + | None => NONEV + | Some loc => SOMEV (#loc) + end. + Local Instance oloc_to_val_inj : Inj (=) (=) oloc_to_val. + Proof. intros [|][|]; simpl; congruence. Qed. + Fixpoint is_list xs v : iProp Σ := - (match xs with - | [] => ⌜v = NONEV⌝ - | x :: xs => ∃ l (t : val), ⌜v = SOMEV #l%V⌝ ∗ l ↦{-} (x, t)%V ∗ is_list xs t + (match xs, v with + | [], None => True + | x :: xs, Some l => ∃ t, l ↦{-} (x, oloc_to_val t)%V ∗ is_list xs t + | _, _ => False end)%I. - Lemma is_list_disj xs v : - is_list xs v -∗ is_list xs v ∗ (⌜v = NONEV⌝ ∨ ∃ l (h t : val), ⌜v = SOMEV #l⌝ ∗ l ↦{-} (h, t)%V). + Lemma is_list_dup xs v : + is_list xs v -∗ is_list xs v ∗ match v with + | None => True + | Some l => ∃ h t, l ↦{-} (h, oloc_to_val t)%V + end. Proof. - destruct xs; auto. - iIntros "H"; iDestruct "H" as (l t) "(-> & Hl & Hstack)". + destruct xs, v; simpl; auto; first by iIntros "[]". + iIntros "H"; iDestruct "H" as (t) "(Hl & Hstack)". iDestruct (partial_mapsto_duplicable with "Hl") as "[Hl1 Hl2]". - iSplitR "Hl2"; first by (iExists _, _; iFrame). iRight; auto. - Qed. - - Lemma is_list_unboxed xs v : - is_list xs v -∗ ⌜val_is_unboxed v⌝ ∗ is_list xs v. - Proof. - iIntros "Hlist"; iDestruct (is_list_disj with "Hlist") as "[$ Heq]". - iDestruct "Heq" as "[-> | H]"; first done; by iDestruct "H" as (? ? ?) "[-> ?]". + iSplitR "Hl2"; first by (iExists _; iFrame). by iExists _, _. Qed. Lemma is_list_empty xs : - is_list xs (InjLV #()) -∗ ⌜xs = []⌝. + is_list xs None -∗ ⌜xs = []⌝. Proof. destruct xs; iIntros "Hstack"; auto. - iDestruct "Hstack" as (? ?) "(% & H)"; discriminate. Qed. Lemma is_list_cons xs l h t : l ↦{-} (h, t)%V -∗ - is_list xs (InjRV #l) -∗ + is_list xs (Some l) -∗ ∃ ys, ⌜xs = h :: ys⌝. Proof. destruct xs; first by iIntros "? %". - iIntros "Hl Hstack"; iDestruct "Hstack" as (l' t') "(% & Hl' & Hrest)"; simplify_eq. + iIntros "Hl Hstack"; iDestruct "Hstack" as (t') "(Hl' & Hrest)". iDestruct (partial_mapsto_agree with "Hl Hl'") as "%"; simplify_eq; iExists _; auto. Qed. Definition stack_inv P l := - (∃ v xs, l ↦ v ∗ is_list xs v ∗ P xs)%I. + (∃ v xs, l ↦ oloc_to_val v ∗ is_list xs v ∗ P xs)%I. Definition is_stack_pred P v := (∃ mailbox l, ⌜v = (mailbox, #l)%V⌝ ∗ is_mailbox P mailbox ∗ inv Nstack (stack_inv P l))%I. @@ -321,7 +325,7 @@ Section proofs. wp_alloc l as "Hl". wp_apply mk_mailbox_works ; first done. iIntros (v) "#Hmailbox". iMod (inv_alloc Nstack _ (stack_inv P l) with "[Hl HP]") as "#Hinv". - { by iNext; iExists _, []; iFrame. } + { by iNext; iExists None, []; iFrame. } wp_pures. iModIntro; iApply "HΦ"; iExists _; auto. Qed. @@ -346,24 +350,25 @@ Section proofs. { iNext; iExists _, _; iFrame. } clear xs. iModIntro. - wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CAS _ _ _). + wp_let. wp_alloc l' as "Hl'". wp_pures. wp_bind (CmpXchg _ _ _). iInv Nstack as (list' xs) "(Hl & Hlist & HP)" "Hclose". - iDestruct (is_list_unboxed with "Hlist") as "[>% Hlist]". destruct (decide (list = list')) as [ -> |]. - * wp_cas_suc. + * wp_cmpxchg_suc. { destruct list'; left; done. } iMod (fupd_intro_mask' (⊤ ∖ ↑Nstack) inner_mask) as "Hupd'"; first solve_ndisj. iMod ("Hupd" with "HP") as "[HP HΨ]". iMod "Hupd'" as "_". iMod ("Hclose" with "[Hl Hl' HP Hlist]") as "_". - { iNext; iExists (SOMEV _), (v' :: xs); iFrame; iExists _, _; iFrame; auto. } + { iNext; iExists (Some _), (v' :: xs); iFrame; iExists _; iFrame; auto. } iModIntro. - wp_if. + wp_pures. by iApply ("HΦ" with "HΨ"). - * wp_cas_fail. + * wp_cmpxchg_fail. + { destruct list, list'; simpl; congruence. } + { destruct list'; left; done. } iMod ("Hclose" with "[Hl HP Hlist]"). { iExists _, _; iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HΦ Hupd"). - wp_match. iApply ("HΦ" with "HΨ"). Qed. @@ -399,8 +404,8 @@ Section proofs. - wp_match. wp_bind (Load _). iInv Nstack as (v xs) "(Hl & Hlist & HP)" "Hclose". wp_load. - iDestruct (is_list_disj with "Hlist") as "[Hlist H]". - iDestruct "H" as "[-> | HSome]". + iDestruct (is_list_dup with "Hlist") as "[Hlist H]". + destruct v as [l'|]; last first. * iDestruct (is_list_empty with "Hlist") as %->. iMod (fupd_intro_mask' (⊤ ∖ ↑Nstack) inner_mask) as "Hupd'"; first solve_ndisj. iMod ("Hupd" with "HP") as "[HP HΨ]". @@ -410,7 +415,7 @@ Section proofs. iModIntro. wp_match. iApply ("HΦ" with "HΨ"). - * iDestruct "HSome" as (l' h t) "[-> Hl']". + * iDestruct "H" as (h t) "Hl'". iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. @@ -421,17 +426,17 @@ Section proofs. iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. - wp_pures. wp_bind (CAS _ _ _). + wp_pures. wp_bind (CmpXchg _ _ _). iInv Nstack as (v' xs'') "(Hl & Hlist & HP)" "Hclose". - destruct (decide (v' = (SOMEV #l'))) as [ -> |]. - + wp_cas_suc. + destruct (decide (v' = (Some l'))) as [ -> |]. + + wp_cmpxchg_suc. iDestruct (is_list_cons with "[Hl'] Hlist") as (ys) "%"; first by iExists _. simplify_eq. iMod (fupd_intro_mask' (⊤ ∖ ↑Nstack) inner_mask) as "Hupd'"; first solve_ndisj. iDestruct "Hupd" as "[Hupdcons _]". iMod ("Hupdcons" with "HP") as "[HP HΨ]". iMod "Hupd'" as "_". - iDestruct "Hlist" as (l'' t') "(% & Hl'' & Hlist)"; simplify_eq. + iDestruct "Hlist" as (t') "(Hl'' & Hlist)". iDestruct "Hl''" as (q') "Hl''". iDestruct (mapsto_agree with "Hl' Hl''") as "%"; simplify_eq. iMod ("Hclose" with "[Hlist Hl HP]") as "_". @@ -439,7 +444,7 @@ Section proofs. iModIntro. wp_pures. iApply ("HΦ" with "HΨ"). - + wp_cas_fail. + + wp_cmpxchg_fail. { destruct v'; simpl; congruence. } iMod ("Hclose" with "[Hlist Hl HP]") as "_". { iNext; iExists _, _; iFrame. } iModIntro. diff --git a/theories/hocap/cg_bag.v b/theories/hocap/cg_bag.v index 758a7389639c5851b9ec372a0d0b68747c76ca8f..062227ab2eacc2cbb9308650a92deb4f2c37e4c0 100644 --- a/theories/hocap/cg_bag.v +++ b/theories/hocap/cg_bag.v @@ -38,9 +38,9 @@ Definition popBag : val := λ: "b", release "l";; "v". -Canonical Structure valmultisetC := leibnizC (gmultiset valC). +Canonical Structure valmultisetO := leibnizO (gmultiset valO). Class bagG Σ := BagG -{ bag_bagG :> inG Σ (prodR fracR (agreeR valmultisetC)); +{ bag_bagG :> inG Σ (prodR fracR (agreeR valmultisetO)); lock_bagG :> lockG Σ }. @@ -80,7 +80,7 @@ Section proof. Proof. rewrite /bag_contents. apply bi.wand_intro_r. rewrite -own_op own_valid uPred.discrete_valid. - f_equiv=> /=. rewrite pair_op. + f_equiv=> /=. rewrite -pair_op. by intros [_ ?%agree_op_invL']. Qed. @@ -90,7 +90,7 @@ Section proof. iIntros "[Hb1 Hb2]". iDestruct (bag_contents_agree with "Hb1 Hb2") as %<-. iMod (own_update_2 with "Hb1 Hb2") as "Hb". - { rewrite pair_op frac_op'. + { rewrite -pair_op frac_op'. assert ((1 / 2 + 1 / 2)%Qp = 1%Qp) as -> by apply Qp_div_2. by apply (cmra_update_exclusive (1%Qp, to_agree Y)). } iDestruct "Hb" as "[Hb1 Hb2]". diff --git a/theories/hocap/concurrent_runners.v b/theories/hocap/concurrent_runners.v index fbbbacd001869100fbc7533369dd16c231cecd95..43f7627674039a5addee44f59a7d6d2e2dfbc5a3 100644 --- a/theories/hocap/concurrent_runners.v +++ b/theories/hocap/concurrent_runners.v @@ -14,7 +14,7 @@ Set Default Proof Using "Type". SET_RES v = the result of the task has been computed and it is v FIN v = the task has been completed with the result v *) (* We use this RA to verify the Task.run() method *) -Definition saR := csumR fracR (csumR (prodR fracR (agreeR valC)) (agreeR valC)). +Definition saR := csumR fracR (csumR (prodR fracR (agreeR valO)) (agreeR valO)). Class saG Σ := { sa_inG :> inG Σ saR }. Definition INIT `{saG Σ} γ (q: Qp) := own γ (Cinl q%Qp). Definition SET_RES `{saG Σ} γ (q: Qp) (v: val) := own γ (Cinr (Cinl (q%Qp, to_agree v))). @@ -32,7 +32,7 @@ Qed. Global Instance SET_RES_fractional `{saG Σ} γ v : Fractional (fun q => SET_RES γ q v)%I. Proof. intros p q. rewrite /SET_RES. - rewrite -own_op Cinr_op Cinl_op pair_op agree_idemp. f_equiv. + rewrite -own_op -Cinr_op -Cinl_op -pair_op agree_idemp. f_equiv. Qed. Global Instance SET_RES_as_fractional `{saG Σ} γ q v: AsFractional (SET_RES γ q v) (fun q => SET_RES γ q v)%I q. @@ -65,7 +65,7 @@ Lemma SET_RES_agree `{saG Σ} (γ: gname) (q q': Qp) (v w: val) : Proof. iIntros "Hs1 Hs2". iDestruct (own_valid_2 with "Hs1 Hs2") as %Hfoo. - iPureIntro. rewrite Cinr_op Cinl_op pair_op in Hfoo. + iPureIntro. rewrite -Cinr_op -Cinl_op -pair_op in Hfoo. by destruct Hfoo as [_ ?%agree_op_invL']. Qed. Lemma FIN_agree `{saG Σ} (γ: gname) (v w: val) : @@ -73,7 +73,7 @@ Lemma FIN_agree `{saG Σ} (γ: gname) (v w: val) : Proof. iIntros "Hs1 Hs2". iDestruct (own_valid_2 with "Hs1 Hs2") as %Hfoo. - iPureIntro. rewrite Cinr_op Cinr_op in Hfoo. + iPureIntro. rewrite -Cinr_op -Cinr_op in Hfoo. by apply agree_op_invL'. Qed. Lemma INIT_SET_RES `{saG Σ} (v: val) γ : @@ -189,7 +189,7 @@ Section contents. Ltac solve_proper ::= solve_proper_core ltac:(fun _ => simpl; auto_equiv). Program Definition pre_runner (γ : name Σ b) (P: val → iProp Σ) (Q: val → val → iProp Σ) : - (valC -n> iProp Σ) -n> (valC -n> iProp Σ) := λne R r, + (valO -n> iProp Σ) -n> (valO -n> iProp Σ) := λne R r, (∃ (body bag : val), ⌜r = (body, bag)%V⌝ ∗ bagS b (N.@"bag") (λ x y, ∃ γ γ', isTask (body,x) γ γ' y P Q) γ bag ∗ ▷ ∀ r a: val, □ (R r ∗ P a -∗ WP body r a {{ v, Q a v }}))%I. @@ -200,7 +200,7 @@ Section contents. Proof. unfold pre_runner. solve_contractive. Qed. Definition runner (γ: name Σ b) (P: val → iProp Σ) (Q: val → val → iProp Σ) : - valC -n> iProp Σ := + valO -n> iProp Σ := (fixpoint (pre_runner γ P Q))%I. Lemma runner_unfold γ r P Q : diff --git a/theories/hocap/fg_bag.v b/theories/hocap/fg_bag.v index 7f174c70b675def0850716333c53fd19787c5622..604d494592c24e6c929e5a38c656b95e594b9413 100644 --- a/theories/hocap/fg_bag.v +++ b/theories/hocap/fg_bag.v @@ -34,9 +34,9 @@ Definition popBag : val := rec: "pop" "b" := else "pop" "b" end. -Canonical Structure valmultisetC := leibnizC (gmultiset valC). +Canonical Structure valmultisetO := leibnizO (gmultiset valO). Class bagG Σ := BagG -{ bag_bagG :> inG Σ (prodR fracR (agreeR valmultisetC)); +{ bag_bagG :> inG Σ (prodR fracR (agreeR valmultisetO)); }. (** Generic specification for the bag, using view shifts. *) @@ -49,49 +49,51 @@ Section proof. Lemma rown_duplicate l v : rown l v -∗ rown l v ∗ rown l v. Proof. iDestruct 1 as (q) "[Hl Hl']". iSplitL "Hl"; iExists _; eauto. Qed. - Fixpoint is_list (hd : val) (xs : list val) : iProp Σ := - match xs with - | [] => ⌜hd = NONEV⌝%I - | x::xs => (∃ (l : loc) (tl : val), - ⌜hd = SOMEV #l⌝ ∗ rown l (x, tl) ∗ is_list tl xs)%I + Definition oloc_to_val (ol: option loc) : val := + match ol with + | None => NONEV + | Some loc => SOMEV (#loc) end. + Local Instance oloc_to_val_inj : Inj (=) (=) oloc_to_val. + Proof. intros [|][|]; simpl; congruence. Qed. - Lemma is_list_unboxed hd xs : - is_list hd xs -∗ ⌜val_is_unboxed hd⌝. - Proof. - destruct xs. - - iIntros (->). done. - - iIntros "Hl". iDestruct "Hl" as (?? ->) "_". done. - Qed. + Fixpoint is_list (hd : option loc) (xs : list val) : iProp Σ := + match xs, hd with + | [], None => True + | x::xs, Some l => ∃ (tl : option loc), + rown l (x, oloc_to_val tl) ∗ is_list tl xs + | _, _ => False + end%I. Lemma is_list_duplicate hd xs : is_list hd xs -∗ is_list hd xs ∗ is_list hd xs. Proof. iInduction xs as [ | x xs ] "IH" forall (hd); simpl; eauto. - iDestruct 1 as (l tl) "[% [Hro Htl]]"; simplify_eq. + destruct hd; last by auto. + iDestruct 1 as (tl) "[Hro Htl]". rewrite rown_duplicate. iDestruct "Hro" as "[Hro Hro']". iDestruct ("IH" with "Htl") as "[Htl Htl']". - iSplitL "Hro Htl"; iExists _,_; iFrame; eauto. + iSplitL "Hro Htl"; iExists _; iFrame; eauto. Qed. Lemma is_list_agree hd xs ys : is_list hd xs -∗ is_list hd ys -∗ ⌜xs = ys⌝. Proof. iInduction xs as [ | x xs ] "IH" forall (hd ys); simpl; eauto. - - iIntros "%"; subst. - destruct ys; eauto. simpl. - iDestruct 1 as (? ?) "[% ?]". simplify_eq. - - iDestruct 1 as (l tl) "(% & Hro & Hls)"; simplify_eq. + - destruct hd; first by auto. + destruct ys; eauto. + - destruct hd; last by auto. destruct ys as [| y ys]; eauto. simpl. - iDestruct 1 as (l' tl') "(% & Hro' & Hls')"; simplify_eq. + iDestruct 1 as (tl) "(Hro & Hls)". + iDestruct 1 as (tl') "(Hro' & Hls')". iDestruct "Hro" as (q) "Hro". iDestruct "Hro'" as (q') "Hro'". - iDestruct (mapsto_agree l' q q' (PairV x tl) (PairV y tl') + iDestruct (mapsto_agree l q q' (PairV x (oloc_to_val tl)) (PairV y (oloc_to_val tl')) with "Hro Hro'") as %?. simplify_eq/=. iDestruct ("IH" with "Hls Hls'") as %->. done. Qed. Definition bag_inv (γb : gname) (b : loc) : iProp Σ := - (∃ (hd : val) (ls : list val), - b ↦ hd ∗ is_list hd ls ∗ own γb ((1/2)%Qp, to_agree (list_to_set_disj ls)))%I. + (∃ (hd : option loc) (ls : list val), + b ↦ oloc_to_val hd ∗ is_list hd ls ∗ own γb ((1/2)%Qp, to_agree (list_to_set_disj ls)))%I. Definition is_bag (γb : gname) (x : val) := (∃ (b : loc), ⌜x = #b⌝ ∗ inv N (bag_inv γb b))%I. Definition bag_contents (γb : gname) (X : gmultiset val) : iProp Σ := @@ -107,7 +109,7 @@ Section proof. Proof. rewrite /bag_contents. apply bi.wand_intro_r. rewrite -own_op own_valid uPred.discrete_valid. - f_equiv=> /=. rewrite pair_op. + f_equiv=> /=. rewrite -pair_op. by intros [_ ?%agree_op_invL']. Qed. @@ -117,7 +119,7 @@ Section proof. iIntros "[Hb1 Hb2]". iDestruct (bag_contents_agree with "Hb1 Hb2") as %<-. iMod (own_update_2 with "Hb1 Hb2") as "Hb". - { rewrite pair_op frac_op'. + { rewrite -pair_op frac_op'. assert ((1 / 2 + 1 / 2)%Qp = 1%Qp) as -> by apply Qp_div_2. by apply (cmra_update_exclusive (1%Qp, to_agree Y)). } iDestruct "Hb" as "[Hb1 Hb2]". @@ -134,7 +136,7 @@ Section proof. wp_alloc r as "Hr". iMod (own_alloc (1%Qp, to_agree ∅)) as (γb) "[Ha Hf]"; first done. iMod (inv_alloc N _ (bag_inv γb r) with "[Ha Hr]") as "#Hinv". - { iNext. iExists _,[]. simpl. iFrame. eauto. } + { iNext. iExists None,[]. simpl. iFrame. } iModIntro. iApply "HΦ". rewrite /is_bag /bag_contents. iFrame. iExists _. by iFrame "Hinv". @@ -160,20 +162,21 @@ Section proof. { iNext. iExists _,_. iFrame. } clear ls. iModIntro. wp_alloc n as "Hn". - wp_pures. wp_bind (CAS _ _ _). + wp_pures. wp_bind (CmpXchg _ _ _). iInv N as (o' ls) "[Ho [Hls >Hb]]" "Hcl". - iPoseProof (is_list_unboxed with "Hls") as "#>%". destruct (decide (o = o')) as [->|?]. - - wp_cas_suc. + - wp_cmpxchg_suc. { destruct o'; left; done. } iMod ("Hvs" with "[$Hb $HP]") as "[Hb HQ]". iMod ("Hcl" with "[Ho Hn Hls Hb]") as "_". - { iNext. iExists _,(v::ls). iFrame "Ho Hb". - simpl. iExists _,_. iFrame. iSplit; eauto. by iExists 1%Qp. } - iModIntro. wp_if_true. by iApply "HΦ". - - wp_cas_fail. + { iNext. iExists (Some _),(v::ls). iFrame "Ho Hb". + simpl. iExists _. iFrame. by iExists 1%Qp. } + iModIntro. wp_pures. by iApply "HΦ". + - wp_cmpxchg_fail. + { destruct o, o'; simpl; congruence. } + { destruct o'; left; done. } iMod ("Hcl" with "[Ho Hls Hb]") as "_". { iNext. iExists _,ls. by iFrame "Ho Hb". } - iModIntro. wp_if_false. + iModIntro. wp_proj. wp_if. by iApply ("IH" with "HP [HΦ]"). Qed. @@ -196,41 +199,41 @@ Section proof. iInv N as (o ls) "[Ho [Hls >Hb]]" "Hcl". wp_load. destruct ls as [|x ls]; simpl. - - iDestruct "Hls" as %->. + - destruct o; first done. iMod ("Hvs2" with "[$Hb $HP]") as "[Hb HQ]". iMod ("Hcl" with "[Ho Hb]") as "_". { iNext. iExists _,[]. by iFrame. } iModIntro. repeat wp_pure _. by iApply "HΦ". - - iDestruct "Hls" as (hd tl) "(% & Hhd & Hls)"; simplify_eq/=. + - destruct o as [hd|]; last done. + iDestruct "Hls" as (tl) "(Hhd & Hls)"; simplify_eq/=. rewrite rown_duplicate. iDestruct "Hhd" as "[Hhd Hhd']". rewrite is_list_duplicate. iDestruct "Hls" as "[Hls Hls']". iMod ("Hcl" with "[Ho Hb Hhd Hls]") as "_". - { iNext. iExists _,(x::ls). simpl; iFrame; eauto. - iExists _, _; eauto. by iFrame. } + { iNext. iExists (Some _),(x::ls). simpl; iFrame; eauto. + iExists _; eauto. by iFrame. } iModIntro. repeat wp_pure _. iDestruct "Hhd'" as (q) "Hhd". - wp_load. repeat wp_pure _. - wp_bind (CAS _ _ _). + wp_load. wp_pures. + wp_bind (CmpXchg _ _ _). iInv N as (o' ls') "[Ho [Hls >Hb]]" "Hcl". - destruct (decide (o' = (InjRV #hd))) as [->|?]. - + wp_cas_suc. + destruct (decide (o' = (Some hd))) as [->|?]. + + wp_cmpxchg_suc. (* The list is still the same *) rewrite (is_list_duplicate tl). iDestruct "Hls'" as "[Hls' Htl]". - iAssert (is_list (InjRV #hd) (x::ls)) with "[Hhd Hls']" as "Hls'". - { simpl. iExists hd,tl. iFrame; iSplit; eauto. - iExists q. iFrame. } + iAssert (is_list (Some hd) (x::ls)) with "[Hhd Hls']" as "Hls'". + { simpl. iExists tl. iFrame. iExists q. iFrame. } iDestruct (is_list_agree with "Hls Hls'") as %?. simplify_eq. iClear "Hls'". - iDestruct "Hls" as (hd' tl') "(% & Hro' & Htl')". simplify_eq. + iDestruct "Hls" as (tl') "(Hro' & Htl')". iMod ("Hvs1" with "[$Hb $HP]") as "[Hb HQ]". iMod ("Hcl" with "[Ho Htl Hb]") as "_". { iNext. iExists _,ls. by iFrame "Ho Hb". } iModIntro. wp_pures. by iApply "HΦ". - + wp_cas_fail. + + wp_cmpxchg_fail. { destruct o'; simpl; congruence. } iMod ("Hcl" with "[Ho Hls Hb]") as "_". { iNext. iExists _,ls'. by iFrame "Ho Hb". } - iModIntro. wp_if_false. + iModIntro. wp_proj. wp_if. by iApply ("IH" with "HP [HΦ]"). Qed. End proof. diff --git a/theories/hocap/lib/oneshot.v b/theories/hocap/lib/oneshot.v index 24a7aa1bf63c6622bce4bb3609ee5d52ce417920..e4d20b6afb4233a04e51be481630ae00bc471de4 100644 --- a/theories/hocap/lib/oneshot.v +++ b/theories/hocap/lib/oneshot.v @@ -6,7 +6,7 @@ Set Default Proof Using "Type". (** We are going to need the oneshot RA to verify the Task.Join() method *) -Definition oneshotR := csumR fracR (agreeR valC). +Definition oneshotR := csumR fracR (agreeR valO). Class oneshotG Σ := { oneshot_inG :> inG Σ oneshotR }. Definition oneshotΣ : gFunctors := #[GFunctor oneshotR]. Instance subG_oneshotΣ {Σ} : subG oneshotΣ Σ → oneshotG Σ. diff --git a/theories/lecture_notes/coq_intro_example_2.v b/theories/lecture_notes/coq_intro_example_2.v index 6e6ced292fa05655d85dc9d1cdeea3c1f993ace4..253a2db8f7ee2b7c9cdc1a8ada88010092af9d30 100644 --- a/theories/lecture_notes/coq_intro_example_2.v +++ b/theories/lecture_notes/coq_intro_example_2.v @@ -59,12 +59,12 @@ Section monotone_counter. *) (* - To tell Coq we wish to use such a discrete CMRA we use the constructor leibnizC. + To tell Coq we wish to use such a discrete CMRA we use the constructor leibnizO. This takes a Coq type and makes it an instance of an OFE (a step-indexed generalization of sets). This is not the place do describe Canonical Structures. A very good introduction is available at https://hal.inria.fr/hal-00816703v1/document *) - Canonical Structure mcounterRAC := leibnizC mcounterRAT. + Canonical Structure mcounterRAC := leibnizO mcounterRAT. (* To make the type mcounterRAT into an RA we need an operation. This is defined in the standard way, except we use the typeclass Op so we can reuse @@ -316,10 +316,10 @@ Section monotone_counter. { iNext; iExists m; iFrame. } iModIntro. wp_let; wp_op; wp_let. - wp_bind (CAS _ _ _)%E. + wp_bind (CmpXchg _ _ _)%E. iInv N as (k) ">[Hpt HOwnAuth]" "HClose". destruct (decide (k = m)); subst. - + wp_cas_suc. + + wp_cmpxchg_suc. (* If the CAS succeeds we need to update our ghost state. This is achieved using the own_update rule/lemma. The arguments are the ghost name and the ghost resources x from which and to which we are updating. Finally we need to give up own γ x to get ownership of the new resources. @@ -332,12 +332,12 @@ Section monotone_counter. iMod ("HClose" with "[Hpt HOwnAuth]") as "_". { iNext; iExists (1 + m)%nat. rewrite Nat2Z.inj_succ Z.add_1_l; iFrame. } - iModIntro; wp_if; iApply ("HCont" with "[HInv HOwnFrag]"). + iModIntro; wp_pures; iApply ("HCont" with "[HInv HOwnFrag]"). iExists γ; iFrame "#"; iFrame. - + wp_cas_fail; first intros ?; simplify_eq. + + wp_cmpxchg_fail; first intros ?; simplify_eq. iMod ("HClose" with "[Hpt HOwnAuth]") as "_". - iExists k; iFrame. - - iModIntro; wp_if. + - iModIntro. wp_proj. wp_if. iApply ("IH" with "HOwnFrag HCont"); iFrame. Qed. End monotone_counter. @@ -501,22 +501,22 @@ Section monotone_counter'. { iNext; iExists m; iFrame. } iModIntro. wp_let; wp_op; wp_let. - wp_bind (CAS _ _ _)%E. + wp_bind (CmpXchg _ _ _)%E. iInv N as (k) ">[Hpt HOwnAuth]" "HClose". destruct (decide (k = m)); subst. - + wp_cas_suc. + + wp_cmpxchg_suc. iMod (own_update γ ((● m ⋅ ◯ n)) (● (S m : mnatUR) ⋅ (◯ (S n : mnatUR))) with "[HOwnFrag HOwnAuth]") as "[HOwnAuth HOwnFrag]". { apply mcounterRA_update'. } { rewrite own_op; iFrame. } iMod ("HClose" with "[Hpt HOwnAuth]") as "_". { iNext; iExists (1 + m)%nat. rewrite Nat2Z.inj_succ Z.add_1_l; iFrame. } - iModIntro; wp_if; iApply ("HCont" with "[HInv HOwnFrag]"). + iModIntro; wp_pures; iApply ("HCont" with "[HInv HOwnFrag]"). iExists γ; iFrame "#"; iFrame. - + wp_cas_fail; first intros ?; simplify_eq. + + wp_cmpxchg_fail; first intros ?; simplify_eq. iMod ("HClose" with "[Hpt HOwnAuth]") as "_". - iExists k; iFrame. - - iModIntro; wp_if. + - iModIntro. wp_proj. wp_if. iApply ("IH" with "HOwnFrag HCont"); iFrame. Qed. End monotone_counter'. @@ -621,16 +621,16 @@ Section ccounter. wp_bind (! _)%E. iInv N as (c) ">[Hγ Hpt]" "Hclose". wp_load. iMod ("Hclose" with "[Hpt Hγ]") as "_"; [iNext; iExists c; by iFrame|]. iModIntro. wp_let. wp_op. wp_let. - wp_bind (CAS _ _ _). iInv N as (c') ">[Hγ Hpt]" "Hclose". + wp_bind (CmpXchg _ _ _). iInv N as (c') ">[Hγ Hpt]" "Hclose". destruct (decide (c' = c)) as [->|]. - iMod (own_update_2 with "Hγ Hown") as "[Hγ Hown]". { apply ccounterRA_update. } (* We use the update lemma for our RA. *) - wp_cas_suc. iMod ("Hclose" with "[Hpt Hγ]") as "_". + wp_cmpxchg_suc. iMod ("Hclose" with "[Hpt Hγ]") as "_". { iNext. iExists (S c). rewrite Nat2Z.inj_succ Z.add_1_l. by iFrame. } - iModIntro. wp_if. iApply "HΦ". by iFrame "Hinv". - - wp_cas_fail; first (by intros [= ?%Nat2Z.inj]). + iModIntro. wp_pures. iApply "HΦ". by iFrame "Hinv". + - wp_cmpxchg_fail; first (by intros [= ?%Nat2Z.inj]). iMod ("Hclose" with "[Hpt Hγ]") as "_"; [iNext; iExists c'; by iFrame|]. - iModIntro. wp_if. by iApply ("IH" with "[Hown] [HΦ]"); auto. + iModIntro. wp_pures. by iApply ("IH" with "[Hown] [HΦ]"); auto. Qed. Lemma read_contrib_spec γ ℓ q n : diff --git a/theories/lecture_notes/lists_guarded.v b/theories/lecture_notes/lists_guarded.v index 04f333627c72c9244f0e82803deeaf539065a658..82e8c6beba0c50be0f22acffb8518cb9b2853ec8 100644 --- a/theories/lecture_notes/lists_guarded.v +++ b/theories/lecture_notes/lists_guarded.v @@ -46,7 +46,7 @@ Notation iProp := (iProp Σ). (* First we define the is_list representation predicate via a guarded fixed point of the functional is_list_pre. Note the use of the later modality. The - arrows -c> express that the arrow is an arrow in the category of COFE's, + arrows -d> express that the arrow is an arrow in the category of COFE's, i.e., it is a non-expansive function. To fully understand the meaning of this it is necessary to understand the model of Iris. @@ -55,7 +55,7 @@ Notation iProp := (iProp Σ). but in more complex examples the domain of the predicate we are defining will not be a discrete type, and the condition will be meaningful and necessary. *) - Definition is_list_pre (Φ : val -c> list val -c> iProp): val -c> list val -c> iProp := λ hd xs, + Definition is_list_pre (Φ : val -d> list val -d> iProp): val -d> list val -d> iProp := λ hd xs, match xs with [] => ⌜hd = NONEV⌝ | (x::xs) => (∃ (ℓ : loc) (hd' : val), ⌜hd = SOMEV #ℓ⌝ ∗ ℓ ↦ (x,hd') ∗ ▷ Φ hd' xs) diff --git a/theories/lecture_notes/lock.v b/theories/lecture_notes/lock.v index 45450c34d41bf4de77238261aca0a5daf73c7026..9db80af79b106943af89f54707f916638404d0a8 100644 --- a/theories/lecture_notes/lock.v +++ b/theories/lecture_notes/lock.v @@ -94,17 +94,19 @@ Section lock_spec. iIntros (HE φ) "#Hi Hcont"; rewrite /acquire. iLöb as "IH". wp_rec. - wp_bind (CAS _ _ _). + wp_bind (CmpXchg _ _ _). iInv (lockN l) as "[(Hl & HP & Ht)|Hl]" "Hcl". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iApply "Hcont". iFrame. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iApply ("IH" with "[$Hcont]"). Qed. @@ -129,4 +131,4 @@ End lock_spec. Typeclasses Opaque locked. Global Opaque locked. Typeclasses Opaque is_lock. -Global Opaque is_lock. \ No newline at end of file +Global Opaque is_lock. diff --git a/theories/lecture_notes/lock_unary_spec.v b/theories/lecture_notes/lock_unary_spec.v index 85eb9810f469fb0f7aa4f2b8b84f6081145f55ef..3f6dcb93cca7e9dc6603dafd10ff3f24531dba1f 100644 --- a/theories/lecture_notes/lock_unary_spec.v +++ b/theories/lecture_notes/lock_unary_spec.v @@ -109,17 +109,19 @@ Section lock_spec. iIntros (HE φ) "#Hi Hcont"; rewrite /acquire. iLöb as "IH". wp_rec. - wp_bind (CAS _ _ _). + wp_bind (CmpXchg _ _ _). iInv (lockN l) as "[(Hl & HP & Ht)|Hl]" "Hcl". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iApply "Hcont". iFrame. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iApply ("IH" with "[$Hcont]"). Qed. @@ -167,16 +169,18 @@ Section lock_spec. iIntros (HE) "#Hi"; rewrite /acquire. iLöb as "IH". wp_rec. - wp_bind (CAS _ _ _). + wp_bind (CmpXchg _ _ _). iInv (lockN l) as "[(Hl & HP & Ht)|Hl]" "Hcl". - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iFrame. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("Hcl" with "[Hl]") as "_"; first by iRight. iModIntro. + wp_proj. wp_if. iApply "IH". Qed. diff --git a/theories/lecture_notes/modular_incr.v b/theories/lecture_notes/modular_incr.v index 2e56aa0a4e8badef8a6988543139d31ff341431e..dab089e25533bdcd4645ab28420beef770e25a61 100644 --- a/theories/lecture_notes/modular_incr.v +++ b/theories/lecture_notes/modular_incr.v @@ -9,7 +9,7 @@ From iris.bi.lib Require Import fractional. From iris.heap_lang.lib Require Import par. -Definition cntCmra : cmraT := (prodR fracR (agreeR (leibnizC Z))). +Definition cntCmra : cmraT := (prodR fracR (agreeR (leibnizO Z))). Class cntG Σ := CntG { CntG_inG :> inG Σ cntCmra }. Definition cntΣ : gFunctors := #[GFunctor cntCmra ]. @@ -169,21 +169,21 @@ Section cnt_spec. iModIntro. wp_let. wp_op. - wp_bind (CAS _ _ _)%E. + wp_bind (CmpXchg _ _ _)%E. iInv (N .@ "internal") as (m') "[>Hpt >Hown]" "HClose". destruct (decide (m = m')); simplify_eq. - - wp_cas_suc. + - wp_cmpxchg_suc. iMod ("HVS" $! m' with "[Hown HP]") as "[Hown HQ]"; first iFrame. iMod ("HClose" with "[Hpt Hown]") as "_". { iNext; iExists _; iFrame. } iModIntro. - wp_if. + wp_pures. iApply "HCont"; by iFrame. - - wp_cas_fail. + - wp_cmpxchg_fail. iMod ("HClose" with "[Hpt Hown]") as "_". { iNext; iExists _; iFrame. } iModIntro. - wp_if. + wp_pures. iApply ("IH" with "HP HCont"). Qed. diff --git a/theories/logatom/conditional_increment/cinc.v b/theories/logatom/conditional_increment/cinc.v index f44a503d53b13ab1ad0a5acb8589023cf4cbf408..694e4558dafa0a437d37fac7669a14ec4b1e5665 100644 --- a/theories/logatom/conditional_increment/cinc.v +++ b/theories/logatom/conditional_increment/cinc.v @@ -15,101 +15,94 @@ Set Default Proof Using "Type". (* new_counter() := let c = ref (injL 0) in - let f = ref false in - ref (f, c) + ref c *) Definition new_counter : val := λ: <>, - let: "c" := ref (ref (InjL #0)) in - let: "f" := ref #true in - ("f", "c"). - -(* - set_flag(ctr, b) := - ctr.1 <- b - *) -Definition set_flag : val := - λ: "ctr" "b", - Fst "ctr" <- "b". + ref (ref (InjL #0)). (* complete(c, f, x, n, p) := - Resolve CAS(c, x, ref (injL (if !f then n+1 else n))) p (ref ()) ;; () + Resolve CmpXchg(c, x, ref (injL (if !f then n+1 else n))) p (ref ()) ;; () *) Definition complete : val := λ: "c" "f" "x" "n" "p", - Resolve (CAS "c" "x" (ref (InjL (if: !"f" then "n" + #1 else "n")))) "p" (ref #()) ;; #(). + let: "l_ghost" := ref #() in + (* Compare with #true to make this a total operation that never + gets stuck, no matter the value stored in [f]. *) + let: "new_n" := if: !"f" = #true then "n" + #1 else "n" in + Resolve (CmpXchg "c" "x" (ref (InjL "new_n"))) "p" "l_ghost" ;; #(). (* - get(c, f) := + get c := let x = !c in match !x with | injL n => n - | injR (n, p) => complete c f x n p; get(c, f) + | injR (f, n, p) => complete c f x n p; get(c, f) *) Definition get : val := - rec: "get" "ctr" := - let: "f" := Fst "ctr" in - let: "c" := Snd "ctr" in + rec: "get" "c" := let: "x" := !"c" in match: !"x" with InjL "n" => "n" | InjR "args" => - complete "c" "f" "x" (Fst "args") (Snd "args") ;; - "get" "ctr" + let: "f" := Fst (Fst "args") in + let: "n" := Snd (Fst "args") in + let: "p" := Snd "args" in + complete "c" "f" "x" "n" "p" ;; + "get" "c" end. (* - cinc(c, f) := + cinc c f := let x = !c in match !x with | injL n => let p := new proph in - let y := ref (injR (n, p)) in + let y := ref (injR (n, f, p)) in if CAS(c, x, y) then complete(c, f, y, n, p) - else cinc(c, f) - | injR (n, p) => - complete(c, f, x, n, p); - cinc(c, f) + else cinc c f + | injR (f', n', p') => + complete(c, f', x, n', p'); + cinc c f *) Definition cinc : val := - rec: "cinc" "ctr" := - let: "f" := Fst "ctr" in - let: "c" := Snd "ctr" in + rec: "cinc" "c" "f" := let: "x" := !"c" in match: !"x" with InjL "n" => let: "p" := NewProph in - let: "y" := ref (InjR ("n", "p")) in + let: "y" := ref (InjR ("f", "n", "p")) in if: CAS "c" "x" "y" then complete "c" "f" "y" "n" "p" ;; Skip - else "cinc" "ctr" + else "cinc" "c" "f" | InjR "args'" => - complete "c" "f" "x" (Fst "args'") (Snd "args'") ;; - "cinc" "ctr" + let: "f'" := Fst (Fst "args'") in + let: "n'" := Snd (Fst "args'") in + let: "p'" := Snd "args'" in + complete "c" "f'" "x" "n'" "p'" ;; + "cinc" "c" "f" end. (** ** Proof setup *) -Definition flagUR := authR $ optionUR $ exclR boolC. -Definition numUR := authR $ optionUR $ exclR ZC. -Definition tokenUR := exclR unitC. -Definition one_shotUR := csumR (exclR unitC) (agreeR unitC). +Definition numUR := authR $ optionUR $ exclR ZO. +Definition tokenUR := exclR unitO. +Definition one_shotUR := csumR (exclR unitO) (agreeR unitO). Class cincG Σ := ConditionalIncrementG { - cinc_flagG :> inG Σ flagUR; cinc_numG :> inG Σ numUR; cinc_tokenG :> inG Σ tokenUR; cinc_one_shotG :> inG Σ one_shotUR; }. Definition cincΣ : gFunctors := - #[GFunctor flagUR; GFunctor numUR; GFunctor tokenUR; GFunctor one_shotUR]. + #[GFunctor numUR; GFunctor tokenUR; GFunctor one_shotUR]. Instance subG_cincΣ {Σ} : subG cincΣ Σ → cincG Σ. Proof. solve_inG. Qed. Section conditional_counter. - Context {Σ} `{!heapG Σ, !cincG Σ}. + Context {Σ} `{!heapG Σ, !gcG Σ, !cincG Σ}. Context (N : namespace). Local Definition stateN := N .@ "state". @@ -124,13 +117,6 @@ Section conditional_counter. by iDestruct "H" as %[H%Excl_included%leibniz_equiv _]%auth_both_valid. Qed. - Lemma sync_flag_values γ_n (b1 b2 : bool) : - own γ_n (● Excl' b1) -∗ own γ_n (◯ Excl' b2) -∗ ⌜b1 = b2⌝. - Proof. - iIntros "H● H◯". iCombine "H●" "H◯" as "H". iDestruct (own_valid with "H") as "H". - by iDestruct "H" as %[H%Excl_included%leibniz_equiv _]%auth_both_valid. - Qed. - Lemma update_counter_value γ_n (n1 n2 m : Z) : own γ_n (● Excl' n1) -∗ own γ_n (◯ Excl' n2) ==∗ own γ_n (● Excl' m) ∗ own γ_n (◯ Excl' m). Proof. @@ -138,34 +124,26 @@ Section conditional_counter. by apply auth_update, option_local_update, exclusive_local_update. Qed. - Lemma update_flag_value γ_n (b1 b2 b : bool) : - own γ_n (● Excl' b1) -∗ own γ_n (◯ Excl' b2) ==∗ own γ_n (● Excl' b) ∗ own γ_n (◯ Excl' b). - Proof. - iIntros "H● H◯". iCombine "H●" "H◯" as "H". rewrite -own_op. iApply (own_update with "H"). - by apply auth_update, option_local_update, exclusive_local_update. - Qed. - - Definition counter_content (γs : gname * gname) (b : bool) (n : Z) := - (own γs.1 (◯ Excl' b) ∗ own γs.2 (◯ Excl' n))%I. - + Definition counter_content (γs : gname) (n : Z) := + (own γs (◯ Excl' n))%I. (** Definition of the invariant *) Fixpoint val_to_some_loc (vs : list (val * val)) : option loc := match vs with - | (#true , LitV (LitLoc l)) :: _ => Some l + | ((_, #true)%V, LitV (LitLoc l)) :: _ => Some l | _ :: vs => val_to_some_loc vs | _ => None end. Inductive abstract_state : Set := - | injl : Z → abstract_state - | injr : Z → proph_id → abstract_state. + | Quiescent : Z → abstract_state + | Updating : loc → Z → proph_id → abstract_state. Definition state_to_val (s : abstract_state) : val := match s with - | injl n => InjLV #n - | injr n p => InjRV (#n,#p) + | Quiescent n => InjLV #n + | Updating f n p => InjRV (#f,#n,#p) end. Global Instance state_to_val_inj : Inj (=) (=) state_to_val. @@ -183,12 +161,12 @@ Section conditional_counter. Definition accepted_state Q (proph_winner : option loc) (l_ghost_winner : loc) := (l_ghost_winner ↦{1/2} - ∗ match proph_winner with None => True | Some l => ⌜l = l_ghost_winner⌝ ∗ Q end)%I. - (* The same thread then wins the CAS and moves from [accepted] to [done]. + (* The same thread then wins the CmpXchg and moves from [accepted] to [done]. Then, the [γ_t] token guards the transition to take out [Q]. - Remember that the thread winning the CAS might be just helping. The token + Remember that the thread winning the CmpXchg might be just helping. The token is owned by the thread whose request this is. In this state, [l_ghost_winner] serves as a token to make sure that - only the CAS winner can transition to here, and owning half of[l] serves as a + only the CmpXchg winner can transition to here, and owning half of[l] serves as a "location" token to ensure there is no ABA going on. Notice how [counter_inv] owns *more than* half of its [l], which means we know that the [l] there and here cannot be the same. *) @@ -206,45 +184,42 @@ Section conditional_counter. ∨ accepted_state Q (val_to_some_loc vs) l_ghost_winner )) ∨ own γ_s (Cinr $ to_agree ()) ∗ done_state Q l l_ghost_winner γ_t))%I. - Definition pau P Q γs := - (▷ P -∗ ◇ AU << ∀ (b : bool) (n : Z), counter_content γs b n >> @ ⊤∖↑N, ∅ - << counter_content γs b (if b then n + 1 else n), COMM Q >>)%I. + Definition pau P Q γs f := + (▷ P -∗ ◇ AU << ∀ (b : bool) (n : Z), counter_content γs n ∗ gc_mapsto f #b >> @ ⊤∖↑N∖↑gcN, ∅ + << counter_content γs (if b then n + 1 else n) ∗ gc_mapsto f #b, COMM Q >>)%I. - Definition counter_inv γ_b γ_n f c := - (∃ (b : bool) (l : loc) (q : Qp) (s : abstract_state), + Definition counter_inv γ_n c := + (∃ (l : loc) (q : Qp) (s : abstract_state), (* We own *more than* half of [l], which shows that this cannot be the [l] of any [state] protocol in the [done] state. *) - f ↦ #b ∗ c ↦{1/2} #l ∗ l ↦{1/2 + q} (state_to_val s) ∗ - own γ_b (● Excl' b) ∗ + c ↦{1/2} #l ∗ l ↦{1/2 + q} (state_to_val s) ∗ match s with - | injl n => c ↦{1/2} #l ∗ own γ_n (● Excl' n) - | injr n p => + | Quiescent n => c ↦{1/2} #l ∗ own γ_n (● Excl' n) + | Updating f n p => ∃ P Q l_ghost_winner γ_t γ_s, (* There are two pieces of per-[state]-protocol ghost state: - [γ_t] is a token owned by whoever created this protocol and used to get out the [Q] in the end. - [γ_s] reflects whether the protocol is [done] yet or not. *) inv stateN (state_inv P Q p n c l l_ghost_winner γ_n γ_t γ_s) ∗ - □ pau P Q (γ_b, γ_n) + □ pau P Q γ_n f ∗ is_gc_loc f end)%I. - Local Hint Extern 0 (environments.envs_entails _ (counter_inv _ _ _ _)) => unfold counter_inv. + Local Hint Extern 0 (environments.envs_entails _ (counter_inv _ _)) => unfold counter_inv. - Definition is_counter (γs : gname * gname) (ctr : val) := - (∃ (γ_b γ_n : gname) (f c : loc), - ⌜γs = (γ_b, γ_n) ∧ ctr = (#f, #c)%V⌝ ∗ - inv counterN (counter_inv γ_b γ_n f c))%I. + Definition is_counter (γ_n : gname) (ctr : val) := + (∃ (c : loc), ⌜ctr = #c ∧ N ## gcN⌝ ∗ gc_inv ∗ inv counterN (counter_inv γ_n c))%I. Global Instance is_counter_persistent γs ctr : Persistent (is_counter γs ctr) := _. - Global Instance counter_content_timeless γs b n : Timeless (counter_content γs b n) := _. + Global Instance counter_content_timeless γs n : Timeless (counter_content γs n) := _. - Global Instance abstract_state_inhabited: Inhabited abstract_state := populate (injl 0). + Global Instance abstract_state_inhabited: Inhabited abstract_state := populate (Quiescent 0). - Lemma counter_content_exclusive γs f1 c1 f2 c2 : - counter_content γs f1 c1 -∗ counter_content γs f2 c2 -∗ False. + Lemma counter_content_exclusive γs c1 c2 : + counter_content γs c1 -∗ counter_content γs c2 -∗ False. Proof. - iIntros "[Hb1 _] [Hb2 _]". iDestruct (own_valid_2 with "Hb1 Hb2") as %?. + iIntros "Hb1 Hb2". iDestruct (own_valid_2 with "Hb1 Hb2") as %?. done. Qed. @@ -281,19 +256,18 @@ Section conditional_counter. (** ** Proof of [complete] *) (** The part of [complete] for the succeeding thread that moves from [accepted] to [done] state *) - Lemma complete_succeeding_thread_pending (γ_b γ_n γ_t γ_s : gname) f_l c_l P Q p + Lemma complete_succeeding_thread_pending (γ_n γ_t γ_s : gname) c_l P Q p (m n : Z) (l l_ghost l_new : loc) Φ : - inv counterN (counter_inv γ_b γ_n f_l c_l) -∗ + inv counterN (counter_inv γ_n c_l) -∗ inv stateN (state_inv P Q p m c_l l l_ghost γ_n γ_t γ_s) -∗ - pau P Q (γ_b, γ_n) -∗ l_ghost ↦{1 / 2} #() -∗ (□(own_token γ_t ={⊤}=∗ ▷ Q) -∗ Φ #()) -∗ own γ_n (● Excl' n) -∗ l_new ↦ InjLV #n -∗ - WP Resolve (CAS #c_l #l #l_new) #p #l_ghost ;; #() {{ v, Φ v }}. + WP Resolve (CmpXchg #c_l #l #l_new) #p #l_ghost ;; #() {{ v, Φ v }}. Proof. - iIntros "#InvC #InvS PAU Hl_ghost HQ Hn● [Hl_new Hl_new']". wp_bind (Resolve _ _ _)%E. - iInv counterN as (b' l' q s) "(>Hf & >Hc & >[Hl Hl'] & Hb● & Hrest)". + iIntros "#InvC #InvS Hl_ghost HQ Hn● [Hl_new Hl_new']". wp_bind (Resolve _ _ _)%E. + iInv counterN as (l' q s) "(>Hc & >[Hl Hl'] & Hrest)". iInv stateN as (vs) "(>Hp & [NotDone | Done])"; last first. { (* We cannot be [done] yet, as we own the "ghost location" that serves as token for that transition. *) @@ -312,9 +286,9 @@ Section conditional_counter. while a [state] protocol is not [done], it owns enough of the [counter] protocol to ensure that does not move anywhere else. *) iDestruct (mapsto_agree with "Hc Hc'") as %[= ->]. - (* We perform the CAS. *) + (* We perform the CmpXchg. *) iCombine "Hc Hc'" as "Hc". - wp_apply (wp_resolve with "Hp"); first done; wp_cas_suc. + wp_apply (wp_resolve with "Hp"); first done; wp_cmpxchg_suc. iIntros (vs' ->) "Hp'". simpl. (* Update to Done. *) iDestruct "Accepted" as "[Hl_ghost_inv [HEq Q]]". @@ -327,32 +301,31 @@ Section conditional_counter. iFrame "#∗". iSplitR "Hl"; iExists _; done. } iModIntro. iSplitR "HQ". { iNext. iDestruct "Hc" as "[Hc1 Hc2]". - iExists _, l_new, _, (injl n). iFrame. } + iExists l_new, _, (Quiescent n). iFrame. } iApply wp_fupd. wp_seq. iApply "HQ". iApply state_done_extract_Q; done. Qed. (** The part of [complete] for the failing thread *) - Lemma complete_failing_thread γ_b γ_n γ_t γ_s f_l c_l l P Q p m n l_ghost_inv l_ghost l_new Φ : + Lemma complete_failing_thread γ_n γ_t γ_s c_l l P Q p m n l_ghost_inv l_ghost l_new Φ : l_ghost_inv ≠ l_ghost → - inv counterN (counter_inv γ_b γ_n f_l c_l) -∗ + inv counterN (counter_inv γ_n c_l) -∗ inv stateN (state_inv P Q p m c_l l l_ghost_inv γ_n γ_t γ_s) -∗ - pau P Q (γ_b, γ_n) -∗ (□(own_token γ_t ={⊤}=∗ ▷ Q) -∗ Φ #()) -∗ l_new ↦ InjLV #n -∗ - WP Resolve (CAS #c_l #l #l_new) #p #l_ghost ;; #() {{ v, Φ v }}. + WP Resolve (CmpXchg #c_l #l #l_new) #p #l_ghost ;; #() {{ v, Φ v }}. Proof. - iIntros (Hnl) "#InvC #InvS PAU HQ Hl_new". wp_bind (Resolve _ _ _)%E. - iInv counterN as (b' l' q s) "(>Hf & >Hc & >[Hl Hl'] & >Hb● & Hrest)". + iIntros (Hnl) "#InvC #InvS HQ Hl_new". wp_bind (Resolve _ _ _)%E. + iInv counterN as (l' q s) "(>Hc & >[Hl Hl'] & Hrest)". iInv stateN as (vs) "(>Hp & [NotDone | [#Hs Done]])". { (* If the [state] protocol is not done yet, we can show that it - is the active protocol still (l = l'). But then the CAS would + is the active protocol still (l = l'). But then the CmpXchg would succeed, and our prophecy would have told us that. So here we can prove that the prophecy was wrong. *) iDestruct "NotDone" as "(_ & >Hc' & State)". iDestruct (mapsto_agree with "Hc Hc'") as %[=->]. iCombine "Hc Hc'" as "Hc". - wp_apply (wp_resolve with "Hp"); first done; wp_cas_suc. + wp_apply (wp_resolve with "Hp"); first done; wp_cmpxchg_suc. iIntros (vs' ->). iDestruct "State" as "[Pending | Accepted]". + iDestruct "Pending" as "[_ [Hvs _]]". iDestruct "Hvs" as %Hvs. by inversion Hvs. + iDestruct "Accepted" as "[_ [Hvs _]]". iDestruct "Hvs" as %Hvs. by inversion Hvs. } @@ -366,11 +339,11 @@ Section conditional_counter. iDestruct (mapsto_combine with "Hl Hl''") as "[Hl _]". rewrite Qp_half_half. iDestruct (mapsto_valid_3 with "Hl Hl'") as "[]". } - (* The CAS fails. *) - wp_apply (wp_resolve with "Hp"); first done. wp_cas_fail. + (* The CmpXchg fails. *) + wp_apply (wp_resolve with "Hp"); first done. wp_cmpxchg_fail. iIntros (vs' ->) "Hp". iModIntro. iSplitL "Done Hp". { iNext. iExists _. iFrame "#∗". } iModIntro. - iSplitL "Hf Hc Hb● Hrest Hl Hl'". { eauto 10 with iFrame. } + iSplitL "Hc Hrest Hl Hl'". { eauto 10 with iFrame. } wp_seq. iApply "HQ". iApply state_done_extract_Q; done. Qed. @@ -380,20 +353,23 @@ Section conditional_counter. this request, then you get [Q]. But we also try to complete other thread's requests, which is why we cannot ask for the token as a precondition. *) - Lemma complete_spec (c f l : loc) (n : Z) (p : proph_id) γ_b γ_n γ_t γ_s l_ghost_inv P Q : - inv counterN (counter_inv γ_b γ_n f c) -∗ + Lemma complete_spec (c f l : loc) (n : Z) (p : proph_id) γ_n γ_t γ_s l_ghost_inv P Q : + N ## gcN → + gc_inv -∗ + inv counterN (counter_inv γ_n c) -∗ inv stateN (state_inv P Q p n c l l_ghost_inv γ_n γ_t γ_s) -∗ - □ pau P Q (γ_b, γ_n) -∗ + □ pau P Q γ_n f -∗ + is_gc_loc f -∗ {{{ True }}} complete #c #f #l #n #p {{{ RET #(); □ (own_token γ_t ={⊤}=∗ ▷Q) }}}. Proof. - iIntros "#InvC #InvS #PAU". + iIntros (?) "#GC #InvC #InvS #PAU #isGC". iModIntro. iIntros (Φ) "_ HQ". wp_lam. wp_pures. - wp_alloc l_ghost as "[Hl_ghost' Hl_ghost'2]". wp_bind (! _)%E. simpl. - (* open outer invariant to read `f` *) - iInv counterN as (b' l' q s) "(>Hf & >Hc & >Hl' & >Hb● & Hrest)". - wp_load. + wp_alloc l_ghost as "[Hl_ghost' Hl_ghost'2]". wp_pures. + wp_bind (! _)%E. simpl. + (* open outer invariant *) + iInv counterN as (l' q s) "(>Hc & >Hl' & Hrest)". (* two different proofs depending on whether we are succeeding thread *) destruct (decide (l_ghost_inv = l_ghost)) as [-> | Hnl]. - (* we are the succeeding thread *) @@ -402,26 +378,29 @@ Section conditional_counter. + (* Pending: update to accepted *) iDestruct "Pending" as "[P >[Hvs Hn●]]". iDestruct ("PAU" with "P") as ">AU". + iMod (gc_access with "GC") as "Hgc"; first solve_ndisj. (* open and *COMMIT* AU, sync flag and counter *) - iMod "AU" as (b2 n2) "[CC [_ Hclose]]". - iDestruct "CC" as "[Hb◯ Hn◯]". simpl. - iDestruct (sync_flag_values with "Hb● Hb◯") as %->. + iMod "AU" as (b n2) "[[Hn◯ Hf] [_ Hclose]]". + iDestruct ("Hgc" with "Hf") as "[Hf Hfclose]". + wp_load. + iMod ("Hfclose" with "Hf") as "[Hf Hfclose]". iDestruct (sync_counter_values with "Hn● Hn◯") as %->. - iMod (update_counter_value _ _ _ (if b2 then n2 + 1 else n2) with "Hn● Hn◯") + iMod (update_counter_value _ _ _ (if b then n2 + 1 else n2) with "Hn● Hn◯") as "[Hn● Hn◯]". - iMod ("Hclose" with "[Hn◯ Hb◯]") as "Q"; first by iFrame. + iMod ("Hclose" with "[Hn◯ Hf]") as "Q"; first by iFrame. + iModIntro. iMod "Hfclose" as "_". (* close state inv *) - iModIntro. iSplitL "Q Hl_ghost' Hp Hvs Hs Hc'". + iIntros "!> !>". iSplitL "Q Hl_ghost' Hp Hvs Hs Hc'". { iNext. iExists _. iFrame "Hp". iLeft. iFrame. iRight. iSplitL "Hl_ghost'"; first by iExists _. destruct (val_to_some_loc vs) eqn:Hvts; iFrame. } (* close outer inv *) iModIntro. iSplitR "Hl_ghost'2 HQ Hn●". { eauto 12 with iFrame. } - destruct b2; wp_if; [ wp_op | .. ]; + destruct b; wp_alloc l_new as "Hl_new"; wp_pures; iApply (complete_succeeding_thread_pending - with "InvC InvS PAU Hl_ghost'2 HQ Hn● Hl_new"). + with "InvC InvS Hl_ghost'2 HQ Hn● Hl_new"). + (* Accepted: contradiction *) iDestruct "Accepted" as "[>Hl_ghost_inv _]". iDestruct "Hl_ghost_inv" as (v) "Hlghost". @@ -432,43 +411,51 @@ Section conditional_counter. iDestruct "Hlghost" as (v) "Hlghost". iCombine "Hl_ghost'" "Hl_ghost'2" as "Hl_ghost'". by iDestruct (mapsto_valid_2 with "Hlghost Hl_ghost'") as %?. - - (* we are the failing thread *) + - (* we are the failing thread. exploit that [f] is a GC location. *) + iMod (is_gc_access with "GC isGC") as (b) "[H↦ Hclose]"; first solve_ndisj. + wp_load. + iMod ("Hclose" with "H↦") as "_". iModIntro. (* close invariant *) - iModIntro. - iSplitL "Hf Hc Hrest Hl' Hb●". { eauto 10 with iFrame. } + iModIntro. iSplitL "Hc Hrest Hl'". { eauto 10 with iFrame. } (* two equal proofs depending on value of b1 *) - destruct b'; wp_if; [ wp_op | ..]; wp_alloc Hl_new as "Hl_new"; wp_pures; + wp_pures. + destruct (bool_decide (b = #true)); + wp_alloc Hl_new as "Hl_new"; wp_pures; iApply (complete_failing_thread - with "InvC InvS PAU HQ Hl_new"); done. + with "InvC InvS HQ Hl_new"); done. Qed. (** ** Proof of [cinc] *) - Lemma cinc_spec γs v : + Lemma cinc_spec γs v (f: loc) : is_counter γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - cinc v @⊤∖↑N - <<< counter_content γs b (if b then n + 1 else n), RET #() >>>. + <<< ∀ (b : bool) (n : Z), counter_content γs n ∗ gc_mapsto f #b >>> + cinc v #f @⊤∖↑N∖↑gcN + <<< counter_content γs (if b then n + 1 else n) ∗ gc_mapsto f #b, RET #() >>>. Proof. - iIntros "#InvC". iDestruct "InvC" as (γ_b γ_n f_l c_l) "[Heq InvC]". - iDestruct "Heq" as %[-> ->]. iIntros (Φ) "AU". iLöb as "IH". - wp_lam. wp_proj. wp_let. wp_proj. wp_let. wp_bind (!_)%E. - iInv counterN as (b' l' q s) "(>Hf & >Hc & >[Hl [Hl' Hl'']] & >Hb● & Hrest)". - wp_load. destruct s as [n|n p]. + iIntros "#InvC". iDestruct "InvC" as (c_l [-> ?]) "[#GC #InvC]". + iIntros (Φ) "AU". iLöb as "IH". + wp_lam. wp_pures. wp_bind (!_)%E. + iInv counterN as (l' q s) "(>Hc & >[Hl [Hl' Hl'']] & Hrest)". + wp_load. destruct s as [n|f' n' p']. - iDestruct "Hrest" as "(Hc' & Hv)". iModIntro. iSplitR "AU Hl'". - { iModIntro. iExists _, _, (q/2)%Qp, (injl n). iFrame. } + { iModIntro. iExists _, (q/2)%Qp, (Quiescent n). iFrame. } wp_let. wp_load. wp_match. wp_apply wp_new_proph; first done. iIntros (l_ghost p') "Hp'". wp_let. wp_alloc ly as "Hly". - wp_let. wp_bind (CAS _ _ _)%E. + wp_let. wp_bind (CmpXchg _ _ _)%E. (* open outer invariant again to CAS c_l *) - iInv counterN as (b2 l'' q2 s) "(>Hf & >Hc & >Hl & >Hb● & Hrest)". + iInv counterN as (l'' q2 s) "(>Hc & >Hl & Hrest)". destruct (decide (l' = l'')) as [<- | Hn]. + (* CAS succeeds *) iDestruct (mapsto_agree with "Hl' Hl") as %<-%state_to_val_inj. iDestruct "Hrest" as ">[Hc' Hn●]". iCombine "Hc Hc'" as "Hc". - wp_cas_suc. + wp_cmpxchg_suc. + (* Take a "peek" at [AU] and abort immediately to get [gc_is_gc f]. *) + iMod "AU" as (b' n') "[[CC Hf] [Hclose _]]". + iDestruct (gc_is_gc with "Hf") as "#Hgc". + iMod ("Hclose" with "[CC Hf]") as "AU"; first by iFrame. (* Initialize new [state] protocol .*) iDestruct (laterable with "AU") as (AU_later) "[AU #AU_back]". iMod (own_alloc (Excl ())) as (γ_t) "Token"; first done. @@ -481,103 +468,83 @@ Section conditional_counter. iFrame. destruct (val_to_some_loc l_ghost); simpl; done. } iModIntro. iDestruct "Hly" as "[Hly1 Hly2]". iSplitR "Hl' Token". { (* close invariant *) - iNext. iExists _, _, _, (injr n p'). eauto 10 with iFrame. + iNext. iExists _, _, (Updating f n p'). eauto 10 with iFrame. } - wp_if. wp_apply complete_spec; [done..|]. + wp_pures. wp_apply complete_spec; [done..|]. iIntros "Ht". iMod ("Ht" with "Token") as "Φ". wp_seq. by wp_lam. + (* CAS fails: closing invariant and invoking IH *) - wp_cas_fail. + wp_cmpxchg_fail. iModIntro. iSplitR "AU". - { iModIntro. iExists _, _, _, s. iFrame. } - wp_if. by iApply "IH". + { iModIntro. iExists _, _, s. iFrame. } + wp_pures. by iApply "IH". - (* l' ↦ injR *) iModIntro. (* extract state invariant *) - iDestruct "Hrest" as (P Q l_ghost γ_t γ_s) "[#InvS #P_AU]". + iDestruct "Hrest" as (P Q l_ghost γ_t γ_s) "(#InvS & #P_AU & #Hgc)". iSplitR "Hl' AU". (* close invariant *) - { iModIntro. iExists _, _, _, (injr n p). iFrame. eauto 10 with iFrame. } + { iModIntro. iExists _, _, (Updating f' n' p'). iFrame. eauto 10 with iFrame. } wp_let. wp_load. wp_match. wp_pures. wp_apply complete_spec; [done..|]. iIntros "_". wp_seq. by iApply "IH". Qed. Lemma new_counter_spec : + N ## gcN → + gc_inv -∗ {{{ True }}} new_counter #() - {{{ ctr γs, RET ctr ; is_counter γs ctr ∗ counter_content γs true 0 }}}. + {{{ ctr γs, RET ctr ; is_counter γs ctr ∗ counter_content γs 0 }}}. Proof. - iIntros (Φ) "_ HΦ". wp_lam. wp_apply wp_fupd. + iIntros (?) "#GC". iIntros (Φ) "!# _ HΦ". wp_lam. wp_apply wp_fupd. wp_alloc l_n as "Hl_n". - wp_alloc l_c as "Hl_c". wp_let. - wp_alloc l_f as "Hl_f". wp_let. wp_pair. - iMod (own_alloc (● Excl' true ⋅ ◯ Excl' true)) as (γ_b) "[Hb● Hb◯]". - { by apply auth_both_valid. } + wp_alloc l_c as "Hl_c". iMod (own_alloc (● Excl' 0 ⋅ ◯ Excl' 0)) as (γ_n) "[Hn● Hn◯]". { by apply auth_both_valid. } - iMod (inv_alloc counterN _ (counter_inv γ_b γ_n l_f l_c) - with "[Hl_f Hl_c Hl_n Hb● Hn●]") as "#InvC". + iMod (inv_alloc counterN _ (counter_inv γ_n l_c) + with "[Hl_c Hl_n Hn●]") as "#InvC". { iNext. iDestruct "Hl_c" as "[Hl_c1 Hl_c2]". iDestruct "Hl_n" as "[??]". - iExists true, l_n, (1/2)%Qp, (injl 0). iFrame. } + iExists l_n, (1/2)%Qp, (Quiescent 0). iFrame. } iModIntro. - iApply ("HΦ" $! (#l_f, #l_c)%V (γ_b, γ_n)). - iSplitR; last by iFrame. iExists γ_b, γ_n, l_f, l_c. iSplit; done. - Qed. - - Lemma set_flag_spec γs v (new_b : bool) : - is_counter γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - set_flag v #new_b @⊤∖↑N - <<< counter_content γs new_b n, RET #() >>>. - Proof. - iIntros "#InvC" (Φ) "AU". iDestruct "InvC" as (γ_b γ_n f_l c_l) "[Heq InvC]". - iDestruct "Heq" as %[-> ->]. wp_lam. wp_let. wp_proj. - iInv counterN as (b c q s) "(>Hf & >Hc & >[Hl Hl'] & >Hb● & Hrest)". - iMod "AU" as (b' n') "[[Hb◯ Hn◯] [_ Hclose]]"; simpl. - wp_store. - iDestruct (sync_flag_values with "Hb● Hb◯") as %HEq; subst b. - iDestruct (update_flag_value with "Hb● Hb◯") as ">[Hb● Hb◯]". - iMod ("Hclose" with "[Hn◯ Hb◯]") as "HΦ"; first by iFrame. - iModIntro. iModIntro. iSplitR "HΦ"; last done. - iNext. iExists new_b, c, q, _. iFrame. done. + iApply ("HΦ" $! #l_c γ_n). + iSplitR; last by iFrame. iExists _. eauto with iFrame. Qed. Lemma get_spec γs v : is_counter γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - get v @⊤∖↑N - <<< counter_content γs b n, RET #n >>>. + <<< ∀ (n : Z), counter_content γs n >>> + get v @⊤∖↑N∖↑gcN + <<< counter_content γs n, RET #n >>>. Proof. - iIntros "#InvC" (Φ) "AU". iDestruct "InvC" as (γ_b γ_n f_l c_l) "[Heq InvC]". - iDestruct "Heq" as %[-> ->]. iLöb as "IH". wp_lam. repeat (wp_proj; wp_let). wp_bind (! _)%E. - iInv counterN as (b c q s) "(>Hf & >Hc & >[Hl [Hl' Hl'']] & >Hb● & Hrest)". + iIntros "#InvC" (Φ) "AU". iDestruct "InvC" as (c_l [-> ?]) "[GC InvC]". + iLöb as "IH". wp_lam. wp_bind (! _)%E. + iInv counterN as (c q s) "(>Hc & >[Hl [Hl' Hl'']] & Hrest)". wp_load. - destruct s as [n|n p]. - - iMod "AU" as (au_b au_n) "[[Hb◯ Hn◯] [_ Hclose]]"; simpl. + destruct s as [n|f n p]. + - iMod "AU" as (au_n) "[Hn◯ [_ Hclose]]"; simpl. iDestruct "Hrest" as "[Hc' Hn●]". iDestruct (sync_counter_values with "Hn● Hn◯") as %->. - iMod ("Hclose" with "[Hn◯ Hb◯]") as "HΦ"; first by iFrame. + iMod ("Hclose" with "Hn◯") as "HΦ". iModIntro. iSplitR "HΦ Hl'". { - iNext. iExists b, c, (q/2)%Qp, (injl au_n). iFrame. + iNext. iExists c, (q/2)%Qp, (Quiescent au_n). iFrame. } wp_let. wp_load. wp_match. iApply "HΦ". - - iDestruct "Hrest" as (P Q l_ghost γ_t γ_s) "[#InvS #PAU]". + - iDestruct "Hrest" as (P Q l_ghost γ_t γ_s) "(#InvS & #PAU & #Hgc)". iModIntro. iSplitR "AU Hl'". { - iNext. iExists b, c, (q/2)%Qp, (injr n p). eauto 10 with iFrame. + iNext. iExists c, (q/2)%Qp, (Updating _ _ p). eauto 10 with iFrame. } - wp_let. wp_load. wp_match. repeat wp_proj. wp_bind (complete _ _ _ _ _)%E. + wp_let. wp_load. wp_pures. wp_bind (complete _ _ _ _ _)%E. wp_apply complete_spec; [done..|]. iIntros "Ht". wp_seq. iApply "IH". iApply "AU". Qed. End conditional_counter. -Definition atomic_cinc `{!heapG Σ, cincG Σ} : +Definition atomic_cinc `{!heapG Σ, !cincG Σ, !gcG Σ} : spec.atomic_cinc Σ := {| spec.new_counter_spec := new_counter_spec; spec.cinc_spec := cinc_spec; - spec.set_flag_spec := set_flag_spec; spec.get_spec := get_spec; spec.counter_content_exclusive := counter_content_exclusive |}. diff --git a/theories/logatom/conditional_increment/spec.v b/theories/logatom/conditional_increment/spec.v index e464ff68850e9ff3ecd0e052464b9748a20e01a2..f816f411edaaf096b632359083f832eaf302a66e 100644 --- a/theories/logatom/conditional_increment/spec.v +++ b/theories/logatom/conditional_increment/spec.v @@ -1,14 +1,14 @@ From stdpp Require Import namespaces. From iris.heap_lang Require Export lifting notation. From iris.program_logic Require Export atomic. +From iris_examples.logatom.lib Require Export gc. Set Default Proof Using "Type". (** A general logically atomic interface for conditional increment. *) -Record atomic_cinc {Σ} `{!heapG Σ} := AtomicCinc { +Record atomic_cinc {Σ} `{!heapG Σ, !gcG Σ} := AtomicCinc { (* -- operations -- *) new_counter : val; cinc : val; - set_flag : val; get : val; (* -- other data -- *) name : Type; @@ -16,34 +16,31 @@ Record atomic_cinc {Σ} `{!heapG Σ} := AtomicCinc { name_countable : Countable name; (* -- predicates -- *) is_counter (N : namespace) (γs : name) (v : val) : iProp Σ; - counter_content (γs : name) (flag : bool) (c : Z) : iProp Σ; + counter_content (γs : name) (c : Z) : iProp Σ; (* -- predicate properties -- *) is_counter_persistent N γs v : Persistent (is_counter N γs v); - counter_content_timeless γs f c : Timeless (counter_content γs f c); - counter_content_exclusive γs f1 c1 f2 c2 : - counter_content γs f1 c1 -∗ counter_content γs f2 c2 -∗ False; + counter_content_timeless γs c : Timeless (counter_content γs c); + counter_content_exclusive γs c1 c2 : + counter_content γs c1 -∗ counter_content γs c2 -∗ False; (* -- operation specs -- *) new_counter_spec N : + N ## gcN → + gc_inv -∗ {{{ True }}} new_counter #() - {{{ ctr γs, RET ctr ; is_counter N γs ctr ∗ counter_content γs true 0 }}}; - cinc_spec N γs v : + {{{ ctr γs, RET ctr ; is_counter N γs ctr ∗ counter_content γs 0 }}}; + cinc_spec N γs v (f : loc) : is_counter N γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - cinc v @⊤∖↑N - <<< counter_content γs b (if b then n + 1 else n), RET #() >>>; - set_flag_spec N γs v (new_b: bool) : - is_counter N γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - set_flag v #new_b @⊤∖↑N - <<< counter_content γs new_b n, RET #() >>>; + <<< ∀ (b : bool) (n : Z), counter_content γs n ∗ gc_mapsto f #b >>> + cinc v #f @⊤∖↑N∖↑gcN + <<< counter_content γs (if b then n + 1 else n) ∗ gc_mapsto f #b, RET #() >>>; get_spec N γs v: is_counter N γs v -∗ - <<< ∀ (b : bool) (n : Z), counter_content γs b n >>> - get v @⊤∖↑N - <<< counter_content γs b n, RET #n >>>; + <<< ∀ (n : Z), counter_content γs n >>> + get v @⊤∖↑N∖↑gcN + <<< counter_content γs n, RET #n >>>; }. -Arguments atomic_cinc _ {_}. +Arguments atomic_cinc _ {_ _}. Existing Instances is_counter_persistent counter_content_timeless diff --git a/theories/logatom/elimination_stack/hocap_spec.v b/theories/logatom/elimination_stack/hocap_spec.v index 8db98bffcffbbff9932918380f554b7d4b0740ab..cda7ac91a3b0585b288fbfa7e7a18de803d7b8fe 100644 --- a/theories/logatom/elimination_stack/hocap_spec.v +++ b/theories/logatom/elimination_stack/hocap_spec.v @@ -138,10 +138,10 @@ auth invariant. *) (** The CMRA & functor we need. *) Class hocapG Σ := HocapG { - hocap_stateG :> inG Σ (authR (optionUR $ exclR (listC valC))); + hocap_stateG :> inG Σ (authR (optionUR $ exclR (listO valO))); }. Definition hocapΣ : gFunctors := - #[GFunctor (exclR unitC); GFunctor (authR (optionUR $ exclR (listC valC)))]. + #[GFunctor (exclR unitO); GFunctor (authR (optionUR $ exclR (listO valO)))]. Instance subG_hocapΣ {Σ} : subG hocapΣ Σ → hocapG Σ. Proof. solve_inG. Qed. diff --git a/theories/logatom/elimination_stack/stack.v b/theories/logatom/elimination_stack/stack.v index 0c8562243d6502774ef8b19149ad6f42e24791c9..1df2057cc8392139774ce43c5afa4900d12600dd 100644 --- a/theories/logatom/elimination_stack/stack.v +++ b/theories/logatom/elimination_stack/stack.v @@ -13,11 +13,11 @@ heap. *) (** The CMRA & functor we need. *) (* Not bundling heapG, as it may be shared with other users. *) Class stackG Σ := StackG { - stack_tokG :> inG Σ (exclR unitC); - stack_stateG :> inG Σ (authR (optionUR $ exclR (listC valC))); + stack_tokG :> inG Σ (exclR unitO); + stack_stateG :> inG Σ (authR (optionUR $ exclR (listO valO))); }. Definition stackΣ : gFunctors := - #[GFunctor (exclR unitC); GFunctor (authR (optionUR $ exclR (listC valC)))]. + #[GFunctor (exclR unitO); GFunctor (authR (optionUR $ exclR (listO valO)))]. Instance subG_stackΣ {Σ} : subG stackΣ Σ → stackG Σ. Proof. solve_inG. Qed. @@ -308,8 +308,7 @@ Section stack. %[->%Excl_included%leibniz_equiv _]%auth_both_valid. destruct l as [|v' l]; simpl. { (* Contradiction. *) iDestruct "Hlist" as ">%". done. } - iDestruct "Hlist" as (tail' q' rep') "[>Heq [>Htail' Hlist]]". - iDestruct "Heq" as %[= <-]. + iDestruct "Hlist" as (tail' q' rep') "[>% [>Htail' Hlist]]". simplify_eq. iDestruct (mapsto_agree with "Htail Htail'") as %[= <- <-%stack_elem_to_val_inj]. iMod (own_update_2 with "Hs● Hl'") as "[Hs● Hl']". { eapply auth_update, option_local_update, (exclusive_local_update _ (Excl _)). done. } diff --git a/theories/logatom/flat_combiner/atomic_sync.v b/theories/logatom/flat_combiner/atomic_sync.v index bf1b078ea6911b47eaf4020393acf3fe69961d5b..7ab80106afa585a4d86ab41fb605df689c763691 100644 --- a/theories/logatom/flat_combiner/atomic_sync.v +++ b/theories/logatom/flat_combiner/atomic_sync.v @@ -6,7 +6,7 @@ From iris_examples.logatom.flat_combiner Require Import sync misc. (** The simple syncer spec in [sync.v] implies a logically atomic spec. *) -Definition syncR := prodR fracR (agreeR valC). (* track the local knowledge of ghost state *) +Definition syncR := prodR fracR (agreeR valO). (* track the local knowledge of ghost state *) Class syncG Σ := sync_tokG :> inG Σ syncR. Definition syncΣ : gFunctors := #[GFunctor (constRF syncR)]. @@ -15,8 +15,8 @@ Proof. solve_inG. Qed. Section atomic_sync. Context `{EqDecision A, !heapG Σ, !lockG Σ}. - Canonical AC := leibnizC A. - Context `{!inG Σ (prodR fracR (agreeR AC))}. + Canonical AO := leibnizO A. + Context `{!inG Σ (prodR fracR (agreeR AO))}. (* TODO: Rename and make opaque; the fact that this is a half should not be visible to the user. *) @@ -49,14 +49,14 @@ Section atomic_sync. Proof. iIntros (Hsync g0 ϕ ret) "Hϕ Hret". iMod (own_alloc (((1 / 2)%Qp, to_agree g0) ⋅ ((1 / 2)%Qp, to_agree g0))) as (γ) "[Hg1 Hg2]". - { by rewrite pair_op agree_idemp. } + { by rewrite -pair_op agree_idemp. } iApply (Hsync (∃ g: A, ϕ g ∗ gHalf γ g)%I with "[Hg1 Hϕ]")=>//. { iExists g0. by iFrame. } iNext. iIntros (s) "#Hsyncer". iApply "Hret". iSplitL "Hg2"; first done. iIntros "!#". iIntros (f). iApply wp_wand_r. iSplitR; first by iApply "Hsyncer". iIntros (f') "#Hsynced {Hsyncer}". - iAlways. iIntros (α β x) "#Hseq". change (ofe_car AC) with A. + iAlways. iIntros (α β x) "#Hseq". change (ofe_car AO) with A. iIntros (Φ') "?". (* TODO: Why can't I iApply "Hsynced"? *) iSpecialize ("Hsynced" $! _ Φ' x). @@ -77,7 +77,7 @@ Section atomic_sync. iAssert (|==> own γ (((1 / 2)%Qp, to_agree g') ⋅ ((1 / 2)%Qp, to_agree g')))%I with "[Hg]" as ">[Hg1 Hg2]". { iApply own_update; last by iAssumption. - apply cmra_update_exclusive. by rewrite pair_op agree_idemp. } + apply cmra_update_exclusive. by rewrite -pair_op agree_idemp. } iMod ("Hvs2" with "[Hg1 Hβ]"). { iExists g'. iFrame. } iModIntro. iSplitL "Hg2 Hϕ'"; last done. diff --git a/theories/logatom/flat_combiner/flat.v b/theories/logatom/flat_combiner/flat.v index a4dd7247bb327da0e44456c8a2c5356999a2d1c7..fec42de26b8c76405fd71d9863cbd98e2a360888 100644 --- a/theories/logatom/flat_combiner/flat.v +++ b/theories/logatom/flat_combiner/flat.v @@ -46,7 +46,7 @@ Definition mk_flat : val := let: "r" := loop "p" "s" "lk" in "r". -Definition reqR := prodR fracR (agreeR valC). (* request x should be kept same *) +Definition reqR := prodR fracR (agreeR valO). (* request x should be kept same *) Definition toks : Type := gname * gname * gname * gname * gname. (* a bunch of tokens to do state transition *) Class flatG Σ := FlatG { req_G :> inG Σ reqR; diff --git a/theories/logatom/flat_combiner/misc.v b/theories/logatom/flat_combiner/misc.v index b31ea9de467860acdd5d05b3f00f0bda884b7c72..fefa7accea6a0fc774c9e7fbaefc82711869055a 100644 --- a/theories/logatom/flat_combiner/misc.v +++ b/theories/logatom/flat_combiner/misc.v @@ -11,16 +11,16 @@ Import uPred. Section lemmas. Lemma pair_l_frac_op' (p q: Qp) (g g': val): ((p + q)%Qp, to_agree g') ~~> (((p, to_agree g') ⋅ (q, to_agree g'))). - Proof. by rewrite pair_op agree_idemp frac_op'. Qed. + Proof. by rewrite -pair_op agree_idemp frac_op'. Qed. Lemma pair_l_frac_op_1' (g g': val): (1%Qp, to_agree g') ~~> (((1/2)%Qp, to_agree g') ⋅ ((1/2)%Qp, to_agree g')). - Proof. by rewrite pair_op agree_idemp frac_op' Qp_div_2. Qed. + Proof. by rewrite -pair_op agree_idemp frac_op' Qp_div_2. Qed. End lemmas. Section excl. - Context `{!inG Σ (exclR unitC)}. + Context `{!inG Σ (exclR unitO)}. Lemma excl_falso γ Q': own γ (Excl ()) ∗ own γ (Excl ()) ⊢ Q'. Proof. diff --git a/theories/logatom/flat_combiner/peritem.v b/theories/logatom/flat_combiner/peritem.v index 2ddcb838cf2eb8a50cc26b9b891c1adc2361127f..44ee033ddd1b44dce1a2efeeeed3552cf7e5d551 100644 --- a/theories/logatom/flat_combiner/peritem.v +++ b/theories/logatom/flat_combiner/peritem.v @@ -70,21 +70,21 @@ Section proofs. wp_load. iMod ("Hclose" with "[Hs H1]"). { iNext. iFrame. iExists xs, hd. iFrame. } iModIntro. wp_let. wp_alloc l as "Hl". - wp_let. wp_bind (CAS _ _ _)%E. + wp_let. wp_bind (CmpXchg _ _ _)%E. iInv N as "H1" "Hclose". iDestruct "H1" as (xs' hd') "[>Hs H1]". destruct (decide (hd = hd')) as [->|Hneq]. - - wp_cas_suc. + - wp_cmpxchg_suc. iMod (inv_alloc N _ (R x) with "[HRx]") as "#HRx"; first eauto. iMod ("Hclose" with "[Hs Hl H1]"). { iNext. iFrame. iExists (x::xs'), l. iFrame. simpl. iExists hd', 1%Qp. iFrame. by iFrame "#". } - iModIntro. wp_if. by iApply "HΦ". - - wp_cas_fail. + iModIntro. wp_pures. by iApply "HΦ". + - wp_cmpxchg_fail. iMod ("Hclose" with "[Hs H1]"). { iNext. iFrame. iExists (xs'), hd'. iFrame. } - iModIntro. wp_if. iApply ("IH" with "HRx"). + iModIntro. wp_pures. iApply ("IH" with "HRx"). by iNext. Qed. diff --git a/theories/logatom/herlihy_wing_queue/hwq.v b/theories/logatom/herlihy_wing_queue/hwq.v new file mode 100644 index 0000000000000000000000000000000000000000..11ae3e5416896f1a1bd749a2dc9ea8b512e6104b --- /dev/null +++ b/theories/logatom/herlihy_wing_queue/hwq.v @@ -0,0 +1,2766 @@ +From iris.algebra Require Import excl auth list gset gmap agree csum. +From iris.heap_lang Require Export lifting notation. +From iris.heap_lang.lib Require Import arith diverge. +From iris.base_logic.lib Require Export invariants proph_map saved_prop. +From iris.program_logic Require Export atomic. +From iris.proofmode Require Import tactics. +From iris.heap_lang Require Import proofmode notation par. +From iris.bi.lib Require Import fractional. +From iris_examples.logatom.herlihy_wing_queue Require Import spec. +Set Default Proof Using "Type". + +(** * Some array-related notations ******************************************) + +Notation "new_array: sz" := + (AllocN sz%E NONEV) (at level 80) : expr_scope. + +Notation "e1 <[[ e2 ]]>" := + (BinOp OffsetOp e1%E e2%E) (at level 8) : expr_scope. + +(** * Implementation of the queue operations ********************************) + +(** new_queue(sz){ + let ar := new_array sz None in + let back := ref 0 in + let p := new_proph () in + {sz, ar, back, p} + } *) +Definition new_queue : val := + λ: "sz", + let: "ar" := new_array: "sz" in + let: "back" := ref #0 in (* First free cell. *) + let: "p" := NewProph in + ("sz", "ar", "back", "p"). + +(** enqueue(q : queue, x : item){ + let i : int := FAA(q.back, 1) in + if(i < q.size){ + q.items[i] := x + } else { + while true; + } + } *) +Definition enqueue : val := + λ: "q" "x", + let: "q_size" := Fst (Fst (Fst "q")) in + let: "q_ar" := Snd (Fst (Fst "q")) in + let: "q_back" := Snd (Fst "q") in + (* Get the next free index. *) + let: "i" := FAA "q_back" #1 in + (* Check not full, and actually insert. *) + if: "i" < "q_size" then "q_ar"<[["i"]]> <- SOME "x" ;; Skip + else diverge #(). + +(** dequeue(q : queue){ + let range = min(!q.back, q.size) in + let rec dequeue_aux(i) = + if i = 0 { + dequeue(q) + } else { + let j = range - i in + let x = ! q.ar[j] in + if x == null { + dequeue_aux(i-1) + } else { + if resolve (CAS q.ar[j] x null) q.p (j, x) { + v + } else { + dequeue_aux(i-1) + } + } + } + in + dequeue_aux(range) + } *) +Definition dequeue_aux : val := + rec: "loop" "dequeue" "q" "range" "i" := + if: "i" = #0 then "dequeue" "q" else + let: "q_ar" := Snd (Fst (Fst "q")) in + let: "q_p" := Snd "q" in + let: "j" := "range" - "i" in + let: "x" := ! "q_ar"<[["j"]]> in + match: "x" with + NONE => "loop" "dequeue" "q" "range" ("i" - #1) + | SOME "v" => + let: "c" := Resolve (CmpXchg ("q_ar"<[["j"]]>) "x" NONE) "q_p" "j" in + if: Snd "c" then "v" else "loop" "dequeue" "q" "range" ("i" - #1) + end. +Definition dequeue : val := + rec: "dequeue" "q" := + let: "q_size" := Fst (Fst (Fst "q")) in + let: "q_back" := Snd (Fst "q") in + let: "range" := minimum !"q_back" "q_size" in + dequeue_aux "dequeue" "q" "range" "range". + +(** * Definition of the cameras we need for queues **************************) + +Definition prod4R A B C D E := + prodR (prodR (prodR (prodR A B) C) D) E. + +Definition oneshotUR := optionUR $ csumR (exclR unitR) (agreeR unitR). +Definition shot : oneshotUR := Some $ Cinr $ to_agree (). +Definition not_shot : oneshotUR := Some $ Cinl $ Excl (). + +Definition per_slot := + prod4R + (* Unique token for the index. *) + (optionUR $ exclR unitR) + (* The location stored at our index, which always remains the same. *) + (optionUR $ agreeR locO) + (* Possible unique name for the index, only if being helped. *) + (optionUR $ exclR gnameO) + (* One shot witnessing the transition from pending to helped. *) + oneshotUR + (* One shot witnessing the physical writing of the value in the slot. *) + oneshotUR. + +Definition eltsUR := authR $ optionUR $ exclR $ listO locO. +Definition contUR := csumR (exclR unitR) (agreeR (prodO natO natO)). +Definition slotUR := authR $ gmapUR nat per_slot. +Definition backUR := authR mnatUR. + +Class hwqG Σ := + HwqG { + hwq_arG :> inG Σ eltsUR; (** Logical contents of the queue. *) + hwq_contG :> inG Σ contUR; (** One-shot for contradiction states. *) + hwq_slotG :> inG Σ slotUR; (** State data for used array slots. *) + hwq_back :> inG Σ backUR; (** Used to show that back only increases. *) + }. + +Definition hwqΣ : gFunctors := + #[GFunctor eltsUR; GFunctor contUR; GFunctor slotUR; GFunctor backUR]. + +Instance subG_hwqΣ {Σ} : subG hwqΣ Σ → hwqG Σ. +Proof. solve_inG. Qed. + +(** * The specifiaction... **************************************************) + +Section herlihy_wing_queue. + +Context `{!heapG Σ, !savedPropG Σ, !hwqG Σ}. +Context (N : namespace). +Notation iProp := (iProp Σ). +Implicit Types γe γc γs : gname. +Implicit Types sz : nat. +Implicit Types ℓ_ar ℓ_back : loc. +Implicit Types p : proph_id. +Implicit Types v : val. +Implicit Types pvs : list nat. + +(** Operations for the CMRA representing the logical contents of the queue. *) + +Lemma new_elts l : (|==> ∃ γe, own γe (● Excl' l) ∗ own γe (◯ Excl' l))%I. +Proof. + iMod (own_alloc (● Excl' l ⋅ ◯ Excl' l)) as (γe) "[H● H◯]". + - by apply auth_both_valid. + - iModIntro. iExists γe. iFrame. +Qed. + +Lemma sync_elts γe (l1 l2 : list loc) : + own γe (● Excl' l1) -∗ own γe (◯ Excl' l2) -∗ ⌜l1 = l2⌝. +Proof. + iIntros "H● H◯". iCombine "H●" "H◯" as "H". + iDestruct (own_valid with "H") as "H". + by iDestruct "H" as %[H%Excl_included%leibniz_equiv _]%auth_both_valid. +Qed. + +Lemma update_elts γe (l1 l2 l : list loc) : + own γe (● Excl' l1) -∗ own γe (◯ Excl' l2) ==∗ + own γe (● Excl' l) ∗ own γe (◯ Excl' l). +Proof. + iIntros "H● H◯". iCombine "H●" "H◯" as "H". rewrite -own_op. + iApply (own_update with "H"). + by apply auth_update, option_local_update, exclusive_local_update. +Qed. + +(* Fragmental part, made available during atomic updates. *) +Definition hwq_cont γe (elts : list loc) : iProp := + own γe (◯ Excl' elts). + +Lemma hwq_cont_exclusive γe elts1 elts2 : + hwq_cont γe elts1 -∗ hwq_cont γe elts2 -∗ False. +Proof. iIntros "H1 H2". by iDestruct (own_valid_2 with "H1 H2") as %?. Qed. + +(** Operations for the CMRA used to show that back only increases. *) + +Definition back_value γb n := own γb (● (n : mnatUR) : backUR). +Definition back_lower_bound γb n := own γb (◯ (n : mnatUR) : backUR). + +Lemma new_back : (|==> ∃ γb, back_value γb 0%nat)%I. +Proof. + iMod (own_alloc (● (0%nat : mnatUR) : backUR)) as (γb) "H●". + - by rewrite auth_auth_valid. + - by iExists γb. +Qed. + +Lemma back_incr γb n : + (back_value γb n ==∗ back_value γb (S n)%nat)%I. +Proof. + iIntros "H●". iMod (own_update with "H●") as "[$ _]"; last done. + apply auth_update_alloc, (mnat_local_update _ _ (S n)). by lia. +Qed. + +Lemma back_snapshot γb n : + (back_value γb n ==∗ back_value γb n ∗ back_lower_bound γb n)%I. +Proof. + iIntros "H●". rewrite -own_op. iApply (own_update with "H●"). + by apply auth_update_alloc, mnat_local_update. +Qed. + +Lemma back_le γb n1 n2 : + (back_value γb n1 -∗ back_lower_bound γb n2 -∗ ⌜n2 ≤ n1⌝%nat)%I. +Proof. + iIntros "H1 H2". iCombine "H1 H2" as "H". + iDestruct (own_valid with "H") as %Hvalid. iPureIntro. + apply auth_both_valid in Hvalid as [H1 H2]. + by apply mnat_included. +Qed. + +(* Stores a lower bound on the [i2] part of any contradiction that + has arised or may arise in the future. *) +Definition i2_lower_bound γi n := back_value γi n. + +(* Witness that the [i2] part of any (future or not) contradicton is + greater than [n]. *) +Definition no_contra_wit γi n := back_lower_bound γi n. + +Lemma i2_lower_bound_update γi n m : + (n ≤ m)%nat → + (i2_lower_bound γi n ==∗ i2_lower_bound γi m)%I. +Proof. + iIntros (H) "H●". iMod (own_update with "H●") as "[$ _]"; last done. + apply auth_update_alloc, (mnat_local_update _ _ m). by lia. +Qed. + +Lemma i2_lower_bound_snapshot γi n : + (i2_lower_bound γi n ==∗ i2_lower_bound γi n ∗ no_contra_wit γi n)%I. +Proof. + iIntros "H●". rewrite -own_op. iApply (own_update with "H●"). + by apply auth_update_alloc, mnat_local_update. +Qed. + +(** Operations for the one-shot CMRA used for contradiction states. *) + +(** Element for "no contradiction yet". *) +Definition no_contra γc := + own γc (Cinl (Excl ())). + +(** Element witnessing a contradiction [(i1, i2)]. *) +Definition contra γc (i1 i2 : nat) := + own γc (Cinr (to_agree (i1, i2))). + +Lemma new_no_contra : (|==> ∃ γc, no_contra γc)%I. +Proof. by apply own_alloc. Qed. + +Lemma to_contra i1 i2 γc : no_contra γc ==∗ contra γc i1 i2. +Proof. apply own_update. by apply cmra_update_exclusive. Qed. + +Lemma contra_not_no_contra i1 i2 γc : + (no_contra γc -∗ contra γc i1 i2 -∗ False)%I. +Proof. iIntros "HnoC HC". iDestruct (own_valid_2 with "HnoC HC") as %[]. Qed. + +Lemma contra_agree i1 i2 i1' i2' γc : + (contra γc i1 i2 -∗ contra γc i1' i2' -∗ ⌜i1' = i1 ∧ i2' = i2⌝)%I. +Proof. + iIntros "HC HC'". iDestruct (own_valid_2 with "HC HC'") as %H. + iPureIntro. apply agree_op_invL' in H. by inversion H. +Qed. + +Global Instance contra_persistent γc i1 i2 : Persistent (contra γc i1 i2). +Proof. apply own_core_persistent. by rewrite /CoreId. Qed. + +(** Operations for the state data. *) + +Inductive state := + (** Help was requested (element not committed). *) + | Pend : gname → state + (** Help has been provided (element committed). *) + | Help : gname → state + (** The enqueue operation known it has been committed. *) + | Done : state. + +Instance state_inhabited : Inhabited state. +Proof. constructor. refine Done. Qed. + +(** Data associated to each slot. The four components are: + - the location that is being written in the slot, + - a possible name for a stored proposition containing the postcondition + of the atomic update of the enqueue happening for the slot (used only + in case of helping), + - state of the slot, + - [true] if a value was physically written in the slot. *) +Definition slot_data : Type := loc * state * bool. + +Implicit Types slots : gmap nat slot_data. + +Definition update_slot i f slots := + match slots !! i with + | Some d => <[i := f d]> (delete i slots) + | None => slots + end. + +Definition val_of (data : slot_data) : loc := + match data with (l, _, _) => l end. + +Definition state_of (data : slot_data) : state := + match data with (_, s, _) => s end. + +Definition name_of (data : slot_data) : option gname := + match state_of data with Pend γ => Some γ | Help γ => Some γ | _ => None end. + +Definition was_written (data : slot_data) : bool := + match data with (_, _, b) => b end. + +Definition was_committed (data : slot_data) : bool := + match state_of data with Pend _ => false | _ => true end. + +Definition set_written (data : slot_data) : slot_data := + match data with (l, s, _) => (l, s, true) end. + +Definition set_written_and_done (data : slot_data) : slot_data := + match data with (l, _, _) => (l, Done, true) end. + +Definition to_helped (γ : gname) (data : slot_data) : slot_data := + match data with (l, _, w) => (l, Help γ, w) end. + +Definition to_done (data : slot_data) : slot_data := + match data with (l, _, w) => (l, Done, w) end. + +Definition physical_value (data : slot_data) : val := + match data with (l, _, w) => if w then SOMEV #l else NONEV end. + +Lemma val_of_set_written d : val_of (set_written d) = val_of d. +Proof. by destruct d as [[l s] w]. Qed. + +Lemma was_written_set_written d : was_written (set_written d) = true. +Proof. by destruct d as [[l s] w]. Qed. + +Lemma state_of_set_written d : state_of (set_written d) = state_of d. +Proof. by destruct d as [[l s] w]. Qed. + +Definition of_slot_data (data : slot_data) : per_slot := + match data with + | (l, s, w) => + let name := match s with Pend γ => Excl' γ | Help γ => Excl' γ | Done => None end in + let comm := if was_committed data then shot else not_shot in + let wr := if w then shot else not_shot in + (Excl' (), Some (to_agree l), name, comm, wr) + end. + +Lemma of_slot_data_valid d : ✓ of_slot_data d. +Proof. by destruct d as [[l []] []]. Qed. + +(* The (unique) token for slot [i]. *) +Definition slot_token γs i := + own γs (◯ {[i := (Excl' (), None, None, None, None)]} : slotUR). + +(* A witness that the location enqueued in slot [i] is [l]. *) +Definition slot_val_wit γs i l := + own γs (◯ {[i := (None, Some (to_agree l), None, None, None)]} : slotUR). + +(* A witness that the element inserted at slot [i] has been committed. *) +Definition slot_committed_wit γs i := + own γs (◯ {[i := (None, None, None, shot, None)]} : slotUR). + +Definition slot_name_tok γs i γ := + own γs (◯ {[i := (None, None, Excl' γ, None, None)]} : slotUR). + +(* A witness that the element inserted at slot [i] has been written. *) +Definition slot_written_wit γs i := + own γs (◯ {[i := (None, None, None, None, shot)]} : slotUR). + +(* A token proving that the enqueue in slot [i] has not been commited. *) +Definition slot_pending_tok γs i := + own γs (◯ {[i := (None, None, None, not_shot, None)]} : slotUR). + +(* A token proving that no value has been written in slot [i]. *) +Definition slot_writing_tok γs i := + own γs (◯ {[i := (None, None, None, None, not_shot)]} : slotUR). + +(* Initial slot data, with not allocated slots. *) +Lemma new_slots : (|==> ∃ γs, own γs (● ∅))%I. +Proof. + iMod (own_alloc (● ∅ ⋅ ◯ ∅)) as (γs) "[H● _]". + - by apply auth_both_valid. + - iModIntro. iExists γs. iFrame. +Qed. + +(* Allocate a new slot with data [d] at the fresh index [i]. *) +Lemma alloc_slot γs slots (i : nat) (d : slot_data) : + slots !! i = None → + own γs (● (of_slot_data <$> slots) : slotUR) ==∗ + own γs (● (of_slot_data <$> (<[i := d]> slots)) : slotUR) ∗ + own γs (◯ {[i := of_slot_data d]} : slotUR). +Proof. + iIntros (Hi) "H". rewrite -own_op fmap_insert. + iApply (own_update with "H"). apply auth_update_alloc. + apply alloc_singleton_local_update. + - by rewrite lookup_fmap Hi. + - apply of_slot_data_valid. +Qed. + +Lemma alloc_done_slot γs slots i l : + slots !! i = None → + own γs (● (of_slot_data <$> slots) : slotUR) ==∗ + own γs (● (of_slot_data <$> (<[i := (l, Done, false)]> slots)) : slotUR) ∗ + slot_token γs i ∗ + slot_val_wit γs i l ∗ + slot_committed_wit γs i ∗ + slot_writing_tok γs i. +Proof. + iIntros (Hi) "H". iMod (alloc_slot _ _ _ _ Hi with "H") as "[$ Hi]". + repeat rewrite -own_op. repeat rewrite -auth_frag_op. + repeat rewrite -insert_op. repeat rewrite left_id. + by rewrite insert_empty. +Qed. + +Lemma alloc_pend_slot γs slots i l γ : + slots !! i = None → + own γs (● (of_slot_data <$> slots) : slotUR) ==∗ + own γs (● (of_slot_data <$> (<[i := (l, Pend γ, false)]> slots)) : slotUR) ∗ + slot_token γs i ∗ + slot_val_wit γs i l ∗ + slot_pending_tok γs i ∗ + slot_name_tok γs i γ ∗ + slot_writing_tok γs i. +Proof. + iIntros (Hi) "H". iMod (alloc_slot _ _ _ _ Hi with "H") as "[$ Hi]". + repeat rewrite -own_op. repeat rewrite -auth_frag_op. + repeat rewrite -insert_op. repeat rewrite left_id. + by rewrite insert_empty. +Qed. + +Lemma use_val_wit γs slots i l : + (own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_val_wit γs i l -∗ + ⌜val_of <$> slots !! i = Some l⌝)%I. +Proof. + iIntros "H● Hwit". iDestruct (own_valid_2 with "H● Hwit") as %H. + iPureIntro. apply auth_both_valid in H as [H%singleton_included _]. + destruct H as [ps (H1 & H2%option_included)]. rewrite lookup_fmap in H1. + destruct (slots !! i) as [d|]; last by inversion H1. simpl in H1. + inversion_clear H1; rename H into H1. + destruct H2 as [H2|[a [b (H21 & H22 & H23)]]]; first done. simplify_eq. + simpl. destruct b as [[[[b1 b2] b3] b4] b5]. + destruct d as [[dl ds] dw]. + destruct H1 as [[[[_ H1] _] _] _]; simpl in H1. simpl. f_equal. + destruct H23 as [H2|H2]. + - destruct H2 as [[[[_ H2] _] _] _]; simpl in H2. + assert (Some (to_agree l) ≡ Some (to_agree dl)) as H by by transitivity b2. + apply Some_equiv_inj, to_agree_inj in H. done. + - apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [_ H2]; simpl in H2. + assert (Some (to_agree l) ≼ Some (to_agree dl)) as H by set_solver. + apply option_included in H. + destruct H as [H|[a [b (H11 & H12 & H13)]]]; first done. + simplify_eq. destruct H13 as [H|H]. + + by apply to_agree_inj in H. + + by apply to_agree_included in H. +Qed. + +Lemma use_name_tok γs slots i γ : + (own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_name_tok γs i γ -∗ + ⌜name_of <$> slots !! i = Some (Some γ)⌝)%I. +Proof. + iIntros "H● Hwit". iDestruct (own_valid_2 with "H● Hwit") as %H. + iPureIntro. apply auth_both_valid in H as [H%singleton_included _]. + destruct H as [ps (H1 & H2%option_included)]. rewrite lookup_fmap in H1. + destruct (slots !! i) as [d|]; last by inversion H1. simpl in H1. + inversion_clear H1; rename H into H1. + destruct H2 as [H2|[a [b (H21 & H22 & H23)]]]; first done. simplify_eq. + simpl. destruct b as [[[[b1 b2] b3] b4] b5]. + destruct d as [[dl ds] dw]. + destruct H1 as [[[[_ _] H1] _] _]; simpl in H1. simpl. f_equal. + destruct H23 as [H2|H2]. + - destruct H2 as [[[[_ _] H2] _] _]; simpl in H2. + destruct ds as [γ'|γ'|]; rewrite /name_of /=; try f_equal. + + assert (Excl' γ ≡ Excl' γ') as H by by transitivity b3. + inversion H as [x y HH|]. by inversion HH. + + assert (Excl' γ ≡ Excl' γ') as H by by transitivity b3. + inversion H as [x y HH|]. by inversion HH. + + assert (Excl' γ ≡ None) as H by by transitivity b3. + inversion H. + - apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [_ H2]; simpl in H2. + destruct ds as [γ'|γ'|]; rewrite /name_of /=; try f_equal. + + assert (Excl' γ ≼ Excl' γ') as H by set_solver. + by apply Excl_included in H. + + assert (Excl' γ ≼ Excl' γ') as H by set_solver. + by apply Excl_included in H. + + assert (Excl' γ ≼ None) as H by set_solver. + exfalso. apply option_included in H as [H|H]; first done. + destruct H as [a [b (H11 & H12 & H13)]]. by simplify_eq. +Qed. + +Lemma shot_not_equiv_not_shot : shot ≢ not_shot. +Proof. + intros H. rewrite /shot /not_shot in H. + inversion H as [x y HAbsurd|]. inversion HAbsurd. +Qed. + +Lemma shot_not_equiv_not_shot' e : shot ≢ not_shot ⋅ e. +Proof. + intros H. rewrite /shot /not_shot in H. + destruct e as [e|]; first destruct e. + - rewrite -Some_op -Cinl_op in H. + inversion H as [x y Habsurd|]; inversion Habsurd. + - rewrite -Some_op in H. compute in H. + inversion H as [x y HAbsurd|]. inversion HAbsurd. + - inversion H as [x y HAbsurd|]. inversion HAbsurd. + - inversion H as [x y HAbsurd|]. inversion HAbsurd. +Qed. + +Lemma shot_not_included_not_shot : ¬ shot ≼ not_shot. +Proof. + intros H. rewrite /shot /not_shot in H. + apply option_included in H. destruct H as [H|H]; first done. + destruct H as [a [b (H1 & H2 & [H3|H3])]]. + - simplify_eq. by inversion H3. + - simplify_eq. apply csum_included in H3. + destruct H3 as [H3|H3]; first done. destruct H3 as [H3|H3]. + + destruct H3 as [a [b (H1 & H2 & H3)]]. by inversion H1. + + destruct H3 as [a [b (H1 & H2 & H3)]]. by inversion H1. +Qed. + +Lemma use_committed_wit γs slots i : + (own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_committed_wit γs i -∗ + ⌜was_committed <$> slots !! i = Some true⌝)%I. +Proof. + iIntros "H● Hwit". iDestruct (own_valid_2 with "H● Hwit") as %H. + iPureIntro. apply auth_both_valid in H as [H%singleton_included _]. + destruct H as [ps (H1 & H2%option_included)]. rewrite lookup_fmap in H1. + destruct (slots !! i) as [d|]; last by inversion H1. simpl in H1. + inversion_clear H1; rename H into H1. + destruct H2 as [H2|[a [b (H21 & H22 & H23)]]]; first done. simplify_eq. + simpl. destruct b as [[[[b1 b2] b3] b4] b5]. + destruct d as [[dl ds] dw]. + destruct H1 as [[[[_ _] _] H1]]; simpl in H1. f_equal. + destruct (was_committed (dl, ds, dw)); first done. exfalso. + destruct H23 as [H2|H2]. + - destruct H2 as [[[[_ _] _] H2] _]; simpl in H2. + apply shot_not_equiv_not_shot. set_solver. + - apply prod_included in H2 as [H2 _]; simpl in H2. + apply prod_included in H2 as [_ H2]; simpl in H2. + apply shot_not_included_not_shot. set_solver. +Qed. + +Lemma use_written_wit γs slots i : + (own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_written_wit γs i -∗ + ⌜was_written <$> slots !! i = Some true⌝)%I. +Proof. + iIntros "H● Hwit". iDestruct (own_valid_2 with "H● Hwit") as %H. + iPureIntro. apply auth_both_valid in H as [H%singleton_included _]. + destruct H as [ps (H1 & H2%option_included)]. rewrite lookup_fmap in H1. + destruct (slots !! i) as [d|]; last by inversion H1. simpl in H1. + inversion_clear H1; rename H into H1. + destruct H2 as [H2|[a [b (H21 & H22 & H23)]]]; first done. simplify_eq. + simpl. destruct b as [[[[b1 b2] b3] b4] b5]. destruct d as [[dl ds] dw]. + destruct H1 as [[[[_ _] _] _] H1]; simpl in H1. f_equal. + destruct dw; first done. exfalso. + destruct H23 as [H2|H2]. + - destruct H2 as [[[[_ _] _] _] H2]; simpl in H2. + exfalso. apply shot_not_equiv_not_shot. set_solver. + - apply prod_included in H2 as [_ H2]; simpl in H2. + exfalso. apply shot_not_included_not_shot. set_solver. +Qed. + +Lemma use_writing_tok γs i slots : + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_writing_tok γs i ==∗ + own γs (● (of_slot_data <$> update_slot i set_written slots) : slotUR) ∗ + slot_written_wit γs i. +Proof. + iIntros "Hs● Htok". iCombine "Hs● Htok" as "H". rewrite -own_op. + iDestruct (own_valid with "H") as %Hvalid. + iApply (own_update with "H"). + apply auth_both_valid in Hvalid as [H1 H2]. + apply singleton_included in H1 as [e (H1_1 & H1_2)]. + rewrite lookup_fmap in H1_1. + destruct (slots !! i) as [[[l s] w]|] eqn:Hi; last by inversion H1_1. + apply Some_equiv_inj in H1_1. + assert (w = false) as ->. + { destruct w; [ exfalso | done ]. + apply Some_included in H1_2 as [H1_2|H1_2]. + - assert ((None, None, None, None, not_shot) + ≡ of_slot_data (l, s, true)) as Hequiv by by transitivity e. + destruct Hequiv as [[[[_ _] _] _] Hequiv]; simpl in Hequiv. + by apply shot_not_equiv_not_shot. + - destruct H1_2 as [f H1_2]. + assert ((None, None, None, None, not_shot) ⋅ f + ≡ of_slot_data (l, s, true)) as Hequiv by by transitivity e. + destruct Hequiv as [[[[_ _] _] _] Hequiv]; simpl in Hequiv. + by eapply shot_not_equiv_not_shot'. } + rewrite /update_slot Hi insert_delete fmap_insert. + apply auth_update. eapply (singleton_local_update _ i). + { by rewrite lookup_fmap Hi. } + rewrite /set_written. apply prod_local_update; first done. simpl. + by apply option_local_update, exclusive_local_update. +Qed. + +Lemma writing_tok_not_written γs slots i : + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_writing_tok γs i -∗ + ⌜was_written <$> slots !! i = Some false⌝. +Proof. + iIntros "Hs● Htok". iCombine "Hs● Htok" as "H". + iDestruct (own_valid with "H") as %Hvalid%auth_both_valid. + iPureIntro. destruct Hvalid as [H1 H2]. + apply singleton_included in H1 as [e (H1_1 & H1_2)]. + rewrite lookup_fmap in H1_1. + destruct (slots !! i) as [[[l s] w]|]; last by inversion H1_1. + apply Some_equiv_inj in H1_1. simpl. f_equal. destruct w; last done. + exfalso. apply Some_included in H1_2 as [H1_2|H1_2]. + - assert ((None, None, None, None, not_shot) + ≡ of_slot_data (l, s, true)) as Hequiv by by transitivity e. + destruct Hequiv as [[[[_ _] _] _] Hequiv]; simpl in Hequiv. + by apply shot_not_equiv_not_shot. + - destruct H1_2 as [f H1_2]. + assert ((None, None, None, None, not_shot) ⋅ f + ≡ of_slot_data (l, s, true)) as Hequiv by by transitivity e. + destruct Hequiv as [[[[_ _] _] _] Hequiv]; simpl in Hequiv. + by eapply shot_not_equiv_not_shot'. +Qed. + +Lemma None_op {A : cmraT} : (None : optionUR A) ⋅ None = None. +Proof. done. Qed. + +Lemma use_pending_tok γs i γ slots : + state_of <$> slots !! i = Some (Pend γ) → + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_pending_tok γs i ==∗ + own γs (● (of_slot_data <$> update_slot i (to_helped γ) slots) : slotUR) ∗ + slot_committed_wit γs i. +Proof. + iIntros (Hlookup) "Hs● Htok". iCombine "Hs● Htok" as "H". + rewrite -own_op. iDestruct (own_valid with "H") as %Hvalid. + iApply (own_update with "H"). + apply auth_both_valid in Hvalid as [H1 H2]. + apply singleton_included in H1 as [e (H1_1 & H1_2)]. + rewrite lookup_fmap in H1_1. + destruct (slots !! i) as [[[l s] w]|] eqn:Hi; last by inversion H1_1. + simpl in Hlookup. inversion Hlookup; subst s. + rewrite /update_slot Hi insert_delete fmap_insert. + apply auth_update. repeat rewrite pair_op. + eapply (singleton_local_update _ i). { by rewrite lookup_fmap Hi. } + rewrite /to_helped. repeat rewrite None_op. + repeat apply prod_local_update; try done. + by apply option_local_update, exclusive_local_update. +Qed. + +Lemma slot_token_exclusive γs i : + slot_token γs i -∗ slot_token γs i -∗ False. +Proof. + iIntros "H1 H2". iCombine "H1 H2" as "H". + iDestruct (own_valid with "H") as %H. iPureIntro. + move:H =>/auth_frag_valid H. apply singleton_valid in H. + by repeat apply pair_valid in H as [H _]; simpl in H. +Qed. + +Lemma helped_to_done_aux γs i γ slots : + state_of <$> slots !! i = Some (Help γ) → + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_name_tok γs i γ ==∗ + own γs (● (of_slot_data <$> update_slot i to_done slots) : slotUR) ∗ + own γs (◯ {[i := (None, None, None, None, None)]} : slotUR). +Proof. + iIntros (H) "H1 H2". iCombine "H1 H2" as "H". + iDestruct (own_valid with "H") as %Hvalid. rewrite -own_op. + iApply (own_update with "H"). apply auth_update. rewrite /update_slot. + destruct (slots !! i) as [d|] eqn:Hd; last by inversion H. + rewrite insert_delete fmap_insert. eapply singleton_local_update. + { by rewrite lookup_fmap Hd /=. } + destruct d as [[dl ds] dw]. inversion H; subst ds; simpl. + repeat apply prod_local_update; try done. simpl. + apply delete_option_local_update. apply _. +Qed. + +Lemma helped_to_done γs i γ slots : + state_of <$> slots !! i = Some (Help γ) → + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_name_tok γs i γ ==∗ + own γs (● (of_slot_data <$> update_slot i to_done slots) : slotUR). +Proof. + iIntros (H) "H1 H2". by iMod (helped_to_done_aux with "H1 H2") as "[H _]". +Qed. + +Lemma val_wit_from_auth γs i l slots : + val_of <$> slots !! i = Some l → + own γs (● (of_slot_data <$> slots) : slotUR) ==∗ + own γs (● (of_slot_data <$> slots) : slotUR) ∗ + slot_val_wit γs i l. +Proof. + iIntros (H) "H". rewrite -own_op. iApply (own_update with "H"). + apply auth_update_core_id; first apply _. + assert (∃ d, slots !! i = Some d) as [d Hlookup]. + { destruct (slots !! i) as [d|]; inversion H. by exists d. } + apply singleton_included. rewrite lookup_fmap. rewrite Hlookup /=. + exists (of_slot_data d). split; first done. + apply Some_included. right. destruct d as [[dl ds] dw]. simpl. + repeat (apply prod_included; split; simpl); + try by (apply option_included; left). + apply option_included; right. exists (to_agree l), (to_agree dl). + repeat (split; first done). left. + rewrite Hlookup /= in H. by inversion H. +Qed. + +(** * Prophecy abstractions *************************************************) + +Fixpoint proph_data sz (deq : gset nat) (rs : list (val * val)) : list nat := + match rs with + | (PairV _ #true , LitV (LitInt i)) :: rs => + if decide (0 ≤ i < sz) then + let i := Z.to_nat i in + if decide (i ∈ deq) then + [] + else + i :: proph_data sz ({[i]} ∪ deq) rs + else [] + | (PairV _ #false, LitV (LitInt i)) :: rs => + if decide (0 ≤ i < sz) then + proph_data sz deq rs + else + [] + | _ => [] + end. + +(* Wrapper for the Iris [proph] proposition, using our data abstraction. *) +Definition hwq_proph p sz deq pvs := + (∃ rs, proph p rs ∗ ⌜pvs = proph_data sz deq rs⌝)%I. + +Lemma proph_data_deq sz deq rs : ∀ i, i ∈ deq → i ∉ proph_data sz deq rs. +Proof. + revert deq. induction rs as [|[b k] rs IH]; intros deq i Hi. + - apply not_elem_of_nil. + - destruct b; simpl; try by apply not_elem_of_nil. + destruct b2; simpl; try by apply not_elem_of_nil. + destruct l; simpl; try by apply not_elem_of_nil. + destruct b. + + destruct k; simpl; try by apply not_elem_of_nil. + destruct l; simpl; try by apply not_elem_of_nil. + destruct (decide (0 ≤ n < sz)); last by apply not_elem_of_nil. + destruct (decide (Z.to_nat n ∈ deq)); first by apply not_elem_of_nil. + apply not_elem_of_cons. split; first by set_solver. + apply IH. set_solver. + + destruct k; simpl; try by apply not_elem_of_nil. + destruct l; simpl; try by apply not_elem_of_nil. + destruct (decide (0 ≤ n < sz)); last by apply not_elem_of_nil. + apply IH. done. +Qed. + +Lemma proph_data_sz sz deq rs : ∀ i, i ∈ proph_data sz deq rs → (i < sz)%nat. +Proof. + revert deq. induction rs as [|[b k] rs IH]; intros deq i Hi. + - set_solver. + - destruct b; simpl; try by set_solver. + destruct b2; simpl; try by set_solver. + destruct l; simpl; try by set_solver. + destruct b. + + destruct k; simpl; try by set_solver. + destruct l; simpl; try by set_solver. + simpl in Hi. + destruct (decide (0 ≤ n < sz)) as [Hn|Hn]; last by set_solver. + destruct (decide (Z.to_nat n ∈ deq)) as [H|H]; first by set_solver. + apply elem_of_cons in Hi. destruct Hi as [->|Hi]. + * apply Nat2Z.inj_lt. destruct Hn as [Hn1 Hn2]. by rewrite Z2Nat.id. + * by apply (IH _ _ Hi). + + destruct k; simpl; try by set_solver. + destruct l; simpl; try by set_solver. + simpl in Hi. + destruct (decide (0 ≤ n < sz)); last by set_solver. + apply (IH _ _ Hi). +Qed. + +Lemma proph_data_NoDup sz deq rs : + NoDup (proph_data sz deq rs ++ elements deq). +Proof. + revert deq. induction rs as [|[b k] rs IH]; intros deq. + - apply NoDup_elements. + - destruct b; simpl; try by apply NoDup_elements. + destruct b2; simpl; try by apply NoDup_elements. + destruct l; simpl; try by apply NoDup_elements. + destruct b. + + destruct k; simpl; try by apply NoDup_elements. + destruct l; simpl; try by apply NoDup_elements. + destruct (decide (0 ≤ n < sz)) + as [Hn|Hn]; last by apply NoDup_elements. + destruct (decide (Z.to_nat n ∈ deq)) + as [Hn_in_deq|Hn_not_in_deq]; first by apply NoDup_elements. + specialize (IH ({[Z.to_nat n]} ∪ deq)) as H1. + assert (Z.to_nat n ∉ proph_data sz ({[Z.to_nat n]} ∪ deq) rs) as H2. + { apply proph_data_deq. by set_solver. } + apply NoDup_app in H1 as (H1_1 & H1_2 & H1_3). + apply NoDup_app. repeat split_and. + * by apply NoDup_cons. + * intros i Hi. apply elem_of_cons in Hi as [Hi|Hi]; first by set_solver. + intros H_elements%elem_of_elements. + eapply (proph_data_deq sz ({[Z.to_nat n]} ∪ deq) rs i); by set_solver. + * by apply NoDup_elements. + + destruct k; simpl; try by apply NoDup_elements. + destruct l; simpl; try by apply NoDup_elements. + destruct (decide (0 ≤ n < sz)); [ by apply IH | by apply NoDup_elements ]. +Qed. + +Definition block : Type := nat * list nat. +Definition blocks : Type := list block. + +(* A block is valid if it follows the structure described above. *) +Definition block_valid slots (b : block) := + slots !! b.1 = None ∧ + ∀ i, i ∈ b.2 → was_committed <$> (slots !! i) = Some false. + +Fixpoint glue_blocks (b : block) (i : nat) (bs : blocks) : blocks := + match bs with + | [] => [b] + | (j, pends) :: bs => if decide (i = j) then (b.1, b.2 ++ i :: pends) :: bs + else b :: glue_blocks (j, pends) i bs + end. + +Fixpoint flatten_blocks bs : list nat := + match bs with + | [] => [] + | (i, pends) :: bs => i :: pends ++ flatten_blocks bs + end. + +Lemma blocks_elem1 b blocks : + b ∈ blocks → b.1 ∈ flatten_blocks blocks. +Proof. + intros H. induction blocks as [|b' blocks IH]; first by inversion H. + destruct (decide (b' = b)) as [->|Hb_not_b']. + - destruct b as [b_u b_ps]. by apply elem_of_list_here. + - destruct b' as [b'_u b'_bs]. simpl. + apply elem_of_list_further. apply elem_of_app; right. + apply IH. apply elem_of_cons in H as [H|H]; last done. + by rewrite H in Hb_not_b'. +Qed. + +Lemma blocks_elem2 b blocks : + b ∈ blocks → ∀ i, i ∈ b.2 → i ∈ flatten_blocks blocks. +Proof. + intros H. induction blocks as [|b' blocks IH]; first by inversion H. + destruct (decide (b' = b)) as [->|Hb_not_b']. + - destruct b as [b_u b_ps]. intros i Hi. simpl in *. + apply elem_of_list_further. apply elem_of_app. by left. + - destruct b' as [b'_u b'_bs]. simpl. intros i Hi. + apply elem_of_list_further. apply elem_of_app; right. + apply IH; last done. apply elem_of_cons in H as [H|H]; last done. + by rewrite H in Hb_not_b'. +Qed. + +Lemma glue_blocks_valid slots i b_unused b_pendings blocks l γ : + slots !! i = None → + b_unused ≠ i → + NoDup (b_unused :: b_pendings ++ flatten_blocks blocks) → + (∀ b : block, b ∈ (b_unused, b_pendings) :: blocks → block_valid slots b) → + ∀ b, b ∈ glue_blocks (b_unused, b_pendings) i blocks → block_valid (<[i:=(l, Pend γ, false)]> slots) b. +Proof using N heapG0 hwqG0 savedPropG0 Σ. + intros Hi. revert b_unused b_pendings. + induction blocks as [|[b_u b_ps] blocks IH]; + intros b_unused b_pendings Hb_unused_not_i HND Hblocks_valid [b_u' b_ps'] Hb. + - apply Hblocks_valid in Hb as Hvalid. + apply elem_of_list_singleton in Hb. simplify_eq. + destruct Hvalid as (Hvalid1 & Hvalid2). split. + + by rewrite lookup_insert_ne. + + simpl in *. intros k Hk. specialize (Hvalid2 _ Hk) as Hvalid_k. + destruct (decide (k = i)) as [->|Hk_not_i]. + * by rewrite lookup_insert. + * by rewrite lookup_insert_ne. + - simpl in Hb. destruct (decide (i = b_u)) as [->|Hi_not_b_u]. + + apply elem_of_cons in Hb as [Hb|Hb]. + * simplify_eq. + assert ((b_unused, b_pendings) ∈ (b_unused, b_pendings) :: (b_u, b_ps) :: blocks) + as Hvalid%Hblocks_valid by set_solver. + destruct Hvalid as (Hvalid1 & Hvalid2). + assert ((b_u, b_ps) ∈ (b_unused, b_pendings) :: (b_u, b_ps) :: blocks) + as Hvalid'%Hblocks_valid by set_solver. + destruct Hvalid' as (Hvalid1' & Hvalid2'). + split; simpl. + ** by rewrite lookup_insert_ne. + ** intros k Hk. apply elem_of_app in Hk as [Hk|Hk]. + *** assert (k ≠ b_u) as HNEq2. + { apply NoDup_cons in HND as (_ & HND). + apply NoDup_app in HND as (_ & HND & _). apply HND in Hk. + simpl in Hk. by apply not_elem_of_cons in Hk as (Hk & _). } + rewrite lookup_insert_ne; last done. by apply Hvalid2. + *** apply elem_of_cons in Hk as [->|Hk]; first by rewrite lookup_insert. + assert (b_u ≠ k) as HNEq2. + { apply NoDup_cons in HND as (_ & HND). + apply NoDup_app in HND as (_ & _ & HND). simpl in HND. + apply NoDup_cons in HND as (HND & _). + apply not_elem_of_app in HND as (HND & _). + intros ->. apply HND, Hk. } + rewrite lookup_insert_ne; last done. by apply Hvalid2'. + * assert ((b_u', b_ps') ∈ (b_unused, b_pendings) :: (b_u, b_ps) :: blocks) + as Hvalid%Hblocks_valid by set_solver. + destruct Hvalid as (Hvalid1 & Hvalid2). rewrite /block_valid. + assert (b_u ≠ b_u') as HNeq1. + { apply NoDup_cons in HND as (_ & HND). + apply NoDup_app in HND as (_ & _ & HND). simpl in HND. + apply NoDup_cons in HND as (HND & _). intros <-. + apply not_elem_of_app in HND as (_ & HND). apply HND. + by apply blocks_elem1 in Hb. } + rewrite lookup_insert_ne; last done. split; first done. + intros k Hk. simpl in Hk. + assert (b_u ≠ k) as HNeq2. + { apply NoDup_cons in HND as (_ & HND). + apply NoDup_app in HND as (_ & _ & HND). simpl in HND. + apply NoDup_cons in HND as (HND & _). intros <-. + apply not_elem_of_app in HND as (_ & HND). apply HND. + by eapply blocks_elem2 in Hb. } + rewrite lookup_insert_ne; last done. by apply Hvalid2. + + apply elem_of_cons in Hb as [Hb|Hb]. + * simplify_eq. + assert ((b_unused, b_pendings) ∈ (b_unused, b_pendings) :: (b_u, b_ps) :: blocks) + as Hvalid%Hblocks_valid by set_solver. + destruct Hvalid as (Hvalid1 & Hvalid2). split. + ** by rewrite lookup_insert_ne. + ** intros k Hk. simpl in *. + assert (k ≠ i) as HNEq. + { intros ->. apply Hvalid2 in Hk. rewrite Hi in Hk. by inversion Hk. } + rewrite lookup_insert_ne; last done. by apply Hvalid2. + * eapply IH; last done. done. + { apply NoDup_cons in HND as (_ & HND). + by apply NoDup_app in HND as (_ & _ & HND). } + intros b' Hb'. + assert (b' ∈ (b_unused, b_pendings) :: (b_u, b_ps) :: blocks) + as Hb'_valid%Hblocks_valid by set_solver. done. +Qed. + +(* Contradiction status: either there is a contradiction going on with + the given indices, or there is no contradiction. In the latter case + the prophecy has well-formed pending blocks as a suffix. *) +Inductive cont_status := + | WithCont : nat → nat → cont_status + | NoCont : blocks → cont_status. + +Global +Instance cont_status_inhabited : Inhabited cont_status. +Proof. constructor. refine (NoCont []). Qed. + +Lemma initial_block_valid b pvs : + b ∈ map (λ i : nat, (i, [])) pvs → block_valid ∅ b. +Proof. + intros H. induction pvs as [|i pvs IH]. + - by inversion H. + - simpl in H. apply elem_of_cons in H as [->|H]. + + split; first by apply lookup_empty. intros k Hk. by inversion Hk. + + apply IH, H. +Qed. + +Lemma flatten_blocks_initial pvs : + pvs = flatten_blocks (map (λ i : nat, (i, [])) pvs). +Proof. + induction pvs as [|i pvs IH]; first done. + simpl. f_equal. by apply IH. +Qed. + +Lemma flatten_blocks_glue b bs i : + flatten_blocks (b :: bs) = flatten_blocks (glue_blocks b i bs). +Proof. + revert b. + induction bs as [|[b_u' b_ps'] bs IH]; intros [b_u b_ps]; first done. + simpl. destruct (decide (i = b_u')) as [->|HNEq]; simpl. + - by rewrite -app_assoc. + - by rewrite -IH. +Qed. + +Lemma flatten_blocks_mem1 blocks : + ∀b, b ∈ blocks → b.1 ∈ flatten_blocks blocks. +Proof. + intros b Hb. induction blocks as [|[i ps] bs IH]; first by inversion Hb. + apply elem_of_cons in Hb as [->|Hb]; first by apply elem_of_list_here. + simpl. apply elem_of_list_further. apply elem_of_app. right. by apply IH. +Qed. + +Lemma flatten_blocks_mem2 blocks : + ∀b, b ∈ blocks → ∀i, i ∈ b.2 → i ∈ flatten_blocks blocks. +Proof. + intros b Hb. induction blocks as [|[i ps] bs IH]; first by inversion Hb. + intros k Hk. apply elem_of_cons in Hb as [->|Hb]; simpl. + - apply elem_of_list_further. apply elem_of_app. by left. + - apply elem_of_list_further. apply elem_of_app. right. by apply IH. +Qed. + +(** * Some definitions and lemmas about array content manipulation **********) + +Definition array_get slots (deqs : gset nat) i := + match slots !! i with + | None => NONEV + | Some d => if decide (i ∈ deqs) then NONEV + else physical_value d + end. + +Fixpoint array_content n slots deqs := + match n with + | 0%nat => [] + | S n => array_content n slots deqs ++ [array_get slots deqs n] + end. + +Lemma length_array_content sz slots deqs : + length (array_content sz slots deqs) = sz. +Proof. + induction sz as [|sz IH]; first done. + by rewrite /= app_length plus_comm /= IH. +Qed. + +Lemma array_content_lookup sz slots deqs i : + (i < sz)%nat → + array_content sz slots deqs !! i = Some (array_get slots deqs i). +Proof. + intros H. induction sz as [|sz IH]; first lia. + destruct (decide (i = sz)) as [->|Hi_not_sz]; simpl. + - rewrite lookup_app_r length_array_content; last done. + by rewrite Nat.sub_diag /=. + - rewrite lookup_app_l; first (apply IH; by lia). + rewrite length_array_content. lia. +Qed. + +Lemma array_content_empty sz : + array_content sz ∅ ∅ = replicate sz NONEV. +Proof. + induction sz as [|sz IH]; first done. + rewrite replicate_S_end /= IH. done. +Qed. + +Lemma array_content_NONEV sz i d slots deqs : + physical_value d = NONEV → slots !! i = None → i ∉ deqs → + array_content sz (<[i:=d]> slots) deqs = array_content sz slots deqs. +Proof. + intros H1 H2 H3. induction sz as [|sz IH]; first done. + rewrite /= /array_get. destruct (decide (i = sz)) as [->|Hi_not_sz]. + - rewrite lookup_insert H2 decide_False; last done. by rewrite IH H1. + - rewrite lookup_insert_ne; last done. by rewrite IH. +Qed. + +Lemma array_content_is_Some sz i slots deqs : + (i < sz)%nat → + is_Some (array_content sz slots deqs !! i). +Proof. + intros H. apply lookup_lt_is_Some. by rewrite length_array_content. +Qed. + +Lemma array_content_ext sz slots1 slots2 deqs : + (∀ i, (i < sz)%nat → array_get slots1 deqs i = array_get slots2 deqs i) → + array_content sz slots1 deqs = array_content sz slots2 deqs. +Proof. + induction sz as [|sz IH]; intros H; first done. + simpl. rewrite H; last by lia. f_equal. apply IH. + intros i Hi. apply H. by lia. +Qed. + +Lemma array_content_more_deqs sz slots deqs i : + (sz ≤ i)%nat → + array_content sz slots ({[i]} ∪ deqs) = array_content sz slots deqs. +Proof. + intros H. induction sz as [|sz IH]; first done. + rewrite /= IH; last by lia. f_equal. + rewrite /array_get. destruct (slots !! sz) as [d|]; last done. + destruct (decide (sz ∈ deqs)) as [Helem|Hnot_elem]. + - rewrite decide_True; [ done | by set_solver ]. + - rewrite decide_False; [ done | .. ]. + apply not_elem_of_union. split; last done. + apply not_elem_of_singleton. by lia. +Qed. + +Lemma array_content_update_slot_ge sz slots deqs f i : + (sz ≤ i)%nat → + array_content sz slots deqs = array_content sz (update_slot i f slots) deqs. +Proof. + intros H. induction sz as [|sz IH]; first done. + rewrite /= IH; last by lia. f_equal. + rewrite /array_get /update_slot. + destruct (slots !! i) as [d|]; last done. + rewrite insert_delete. rewrite lookup_insert_ne; [ done | by lia ]. +Qed. + +Lemma array_content_dequeue sz i slots deqs : + (i < sz)%nat → + i ∉ deqs → + array_content sz slots ({[i]} ∪ deqs) = <[i:=NONEV]> (array_content sz slots deqs). +Proof using N heapG0 hwqG0 savedPropG0 Σ. + revert i. induction sz as [|sz IH]; intros i H1 H2; first done. + destruct (decide (sz = i)) as [->|Hsz_not_i]; simpl. + - assert (i = length (array_content i slots deqs) + 0)%nat as HEq. + { rewrite length_array_content. by lia. } + rewrite [X in <[X:=_]> _]HEq. + rewrite (insert_app_r (array_content i slots deqs) _ 0 NONEV). + rewrite /= /array_get. destruct (slots !! i) as [d|]. + + rewrite decide_True; last by set_solver. f_equal. + rewrite array_content_more_deqs; [ done | by lia ]. + + f_equal. rewrite array_content_more_deqs; [ done | by lia ]. + - rewrite insert_app_l; last (rewrite length_array_content; by lia). + rewrite IH; [ .. | by lia | done ]. f_equal. + rewrite /array_get. destruct (slots !! sz) as [d|]; last done. + destruct (decide (sz ∈ deqs)) as [H|H]. + * rewrite decide_True; [ done | by set_solver ]. + * rewrite decide_False; [ done | by set_solver ]. +Qed. + +Lemma array_content_set_written sz i (l : loc) slots deqs : + (i < sz)%nat → + val_of <$> slots !! i = Some l → + ¬ i ∈ deqs → + <[i:=InjRV #l]> (array_content sz slots deqs) = array_content sz (update_slot i set_written slots) deqs. +Proof using N heapG0 hwqG0 savedPropG0 Σ. + revert i. induction sz as [|sz IH]; intros i H1 H2 H3; first done. + destruct (decide (sz = i)) as [->|Hsz_not_i]; simpl. + - assert (i = length (array_content i slots deqs) + 0)%nat as HEq. + { rewrite length_array_content. by lia. } + rewrite [X in <[X:=_]> _]HEq. + rewrite (insert_app_r (array_content i slots deqs) _ 0). + erewrite array_content_update_slot_ge; [ f_equal | by lia ]. + rewrite /= /array_get /update_slot. destruct (slots !! i) as [d|]. + + rewrite lookup_insert decide_False; last done. + destruct d as [[ld sd] wd]. inversion H2; subst ld. done. + + inversion H2. + - rewrite insert_app_l; last (rewrite length_array_content; by lia). + rewrite IH; [ .. | by lia | done | done ]. f_equal. + rewrite /array_get /update_slot. destruct (slots !! i) as [d|]; last done. + by rewrite insert_delete lookup_insert_ne. +Qed. + +(* FIXME similar to previous lemma. Share stuff? *) +Lemma array_content_set_written_and_done sz i (l : loc) slots deqs : + (i < sz)%nat → + val_of <$> slots !! i = Some l → + ¬ i ∈ deqs → + <[i:=InjRV #l]> (array_content sz slots deqs) = array_content sz (update_slot i set_written_and_done slots) deqs. +Proof. + revert i. induction sz as [|sz IH]; intros i H1 H2 H3; first done. + destruct (decide (sz = i)) as [->|Hsz_not_i]; simpl. + - assert (i = length (array_content i slots deqs) + 0)%nat as HEq. + { rewrite length_array_content. by lia. } + rewrite [X in <[X:=_]> _]HEq. + rewrite (insert_app_r (array_content i slots deqs) _ 0). + erewrite array_content_update_slot_ge; [ f_equal | by lia ]. + rewrite /= /array_get /update_slot. destruct (slots !! i) as [d|]. + + rewrite lookup_insert decide_False; last done. + destruct d as [[ld sd] wd]. inversion H2; subst ld. done. + + inversion H2. + - rewrite insert_app_l; last (rewrite length_array_content; by lia). + rewrite IH; [ .. | by lia | done | done ]. f_equal. + rewrite /array_get /update_slot. destruct (slots !! i) as [d|]; last done. + by rewrite insert_delete lookup_insert_ne. +Qed. + +Lemma update_slot_lookup i f slots : + update_slot i f slots !! i = f <$> slots !! i. +Proof. + rewrite /update_slot. + destruct (slots !! i) as [d|] eqn:HEq; last done. + by rewrite lookup_insert. +Qed. + +Lemma update_slot_lookup_ne i k f slots : + i ≠ k → + update_slot i f slots !! k = slots !! k. +Proof. + intros H. rewrite /update_slot. + destruct (slots !! i) as [d|] eqn:HEq; last done. + rewrite lookup_insert_ne; last done. + by rewrite lookup_delete_ne. +Qed. + +Lemma update_slot_update_slot i f g slots : + update_slot i f (update_slot i g slots) = update_slot i (f ∘ g) slots. +Proof. + rewrite /update_slot. + destruct (slots !! i) as [d|] eqn:HEq. + - rewrite lookup_insert. repeat rewrite insert_delete. + rewrite insert_insert. done. + - rewrite HEq. done. +Qed. + +Definition get_value slots (deqs : gset nat) i := + match slots !! i with + | None => Build_loc 0 + | Some d => val_of d + end. + +Definition map_get_value_not_in_pref i d pref slots deqs : + was_written d = false → + i ∉ pref → + map (get_value (<[i:=d]> slots) deqs) pref = map (get_value slots deqs) pref. +Proof. + intros Hd. induction pref as [|k pref IH]; intros Hi; first done. + rewrite /= IH; last by set_solver. f_equal. rewrite /get_value. + rewrite lookup_insert_ne; first done. set_solver. +Qed. + +(** * Definition of the main ************************************************) + +(** Atomic update for the insertion of [l], with post-condition [Q]. *) +Definition enqueue_AU γe l Q := + (AU << ∀ ls : list loc, hwq_cont γe ls >> @ ⊤ ∖ ↑N, ∅ + << hwq_cont γe (ls ++ [l]), COMM Q >>)%I. + +(* +When a contradiction is going on, we have [cont = WithCont i1 i2] where: + - [i1] is the index reserved by the enqueue operation the initiated the + contradiction, + - [i2] is the first index in the prophecy that was not yet reserved for + an enqueue operation (when the contradiction was initiated). +*) + +Definition per_slot_own γe γs i d := + (slot_val_wit γs i (val_of d) ∗ + (if was_written d then slot_written_wit γs i else True) ∗ + match state_of d with + | Pend γ => slot_pending_tok γs i ∗ + ∃ Q, saved_prop_own γ Q ∗ enqueue_AU γe (val_of d) Q + | Help γ => slot_committed_wit γs i ∗ ∃ Q, saved_prop_own γ Q ∗ ▷ Q + | Done => slot_committed_wit γs i ∗ slot_token γs i + end)%I. + +Definition inv_hwq sz γb γi γe γc γs ℓ_ar ℓ_back p : iProp := + (∃ (back : nat) (** Physical value of [q.back]. *) + (pvs : list nat) (** Full contents of the prophecy. *) + (pref : list nat) (** Commit prefix of the prophecy *) + (rest : list loc) (** Logical queue after commit prefix. *) + (cont : cont_status) (** Contradiction or prophecy suffix. *) + (slots : gmap nat slot_data) (** Per-slot data for used indices. *) + (deqs : gset nat), (** Dequeued indices. *) + (** Physical data. *) + ℓ_back ↦ #back ∗ ℓ_ar ↦∗ (array_content sz slots deqs) ∗ + (** Logical contents of the queue and prophecy contents. *) + back_value γb back ∗ + i2_lower_bound γi (match cont with WithCont _ i2 => i2 | NoCont _ => back `min` sz end)%nat ∗ + own γe (● (Excl' (map (get_value slots deqs) pref ++ rest))) ∗ + own γs (● (of_slot_data <$> slots : gmap nat per_slot)) ∗ + hwq_proph p sz deqs pvs ∗ + (** Per-slot ownership. *) + ([∗ map] i ↦ d ∈ slots, per_slot_own γe γs i d) ∗ + (** Contradiction status. *) + match cont with NoCont _ => no_contra γc | WithCont i1 i2 => contra γc i1 i2 end ∗ + (** Tying the logical and physical data and some other pure stuff. *) + ⌜(∀ i, (i < back `min` sz)%nat ↔ is_Some (slots !! i)) ∧ + (∀ i, (was_committed <$> slots !! i = Some false → was_written <$> slots !! i = Some false) ∧ + (was_written <$> slots !! i = Some false → i ∉ deqs)) ∧ + (∀ i, i ∈ pref → was_committed <$> slots !! i = Some true ∧ i ∉ deqs ∧ + match cont with WithCont i1 _ => i ≠ i1 | _ => True end) ∧ + (∀ i, i ∈ deqs → was_written <$> slots !! i = Some true ∧ + was_committed <$> slots !! i = Some true ∧ + array_get slots deqs i = NONEV) ∧ + (NoDup (pvs ++ elements deqs) ∧ ∀ i, i ∈ pvs → (i < sz)%nat) ∧ + match cont with + | NoCont bs => + (∀ b, b ∈ bs → block_valid slots b) ∧ + (bs ≠ [] → rest = []) ∧ + pvs = pref ++ flatten_blocks bs + | WithCont i1 i2 => + (i1 < i2 < sz ∧ i1 < back)%nat ∧ + was_committed <$> slots !! i1 = Some true ∧ + was_written <$> slots !! i1 = Some true ∧ ¬ i1 ∈ deqs ∧ + array_get slots deqs i1 ≠ NONEV ∧ + pref ++ [i2] `prefix_of` pvs + end⌝)%I. + +Definition is_hwq sz γe v : iProp := + (∃ γb γi γc γs ℓ_ar ℓ_back p, + ⌜v = (#sz, #ℓ_ar, #ℓ_back, #p)%V⌝ ∗ + inv N (inv_hwq sz γb γi γe γc γs ℓ_ar ℓ_back p))%I. + +(** * Some useful instances *************************************************) + +Instance blocks_match_persistent (bs : blocks) γc i1 : + Persistent (match bs with + | [] => True + | (i2, _) :: _ => contra γc i1 i2 + end)%I. +Proof. destruct bs as [|[i2 _] _]; apply _. Qed. + +Instance cont_match_persistent cont γc : + Persistent (match cont with + | NoCont _ => True + | WithCont i1 i2 => contra γc i1 i2 + end)%I. +Proof. destruct cont as [i1 i2|_]; apply _. Qed. + +Instance contra_timeless cont γc : + Timeless (match cont with + | NoCont _ => no_contra γc + | WithCont i1 i2 => contra γc i1 i2 + end). +Proof. destruct cont as [i1 i2|_]; apply _. Qed. + +(** * Some important lemmas for the specification of [enqueue] **************) + +Definition get_values (slots : gmap nat slot_data) (p : list nat) := + fold_right (λ i acc, match val_of <$> slots !! i with + | None => acc + | Some l => l :: acc end) [] p. + +Definition get_values_not_in n ps d s : + n ∉ ps → get_values (<[n:=d]> s) ps = get_values s ps. +Proof. + intros H. induction ps as [|p ps IH]; first done. simpl. + assert (n ≠ p) as Hn_not_p by set_solver. + rewrite lookup_insert_ne; last done. + rewrite IH; first done. set_solver. +Qed. + +Definition helped (p : list nat) (i : nat) d := + match state_of d with + | Pend γ => if decide (i ∈ p) then + Some (val_of d, Help γ, was_written d) + else + Some d + | _ => Some d + end. + +Lemma is_Some_helped (p : list nat) i d : is_Some (helped p i d). +Proof. + rewrite /helped. destruct (state_of d); try by eexists. + destruct (decide (i ∈ p)); by eexists. +Qed. + +Lemma map_imap_helped_nil slots : map_imap (helped []) slots = slots. +Proof. + apply map_eq. intros i. rewrite map_lookup_imap. + destruct (slots !! i) as [d|] eqn:HEq. + - rewrite /helped /= HEq. by destruct (state_of d). + - by rewrite /= HEq. +Qed. + +Lemma annoying_lemma_1 slots deqs pref i l b_pendings : + (∀ k, k ∈ pref → was_committed <$> slots !! k = Some true ∧ k ∉ deqs) → + NoDup (pref ++ i :: b_pendings) → + map (get_value (map_imap (helped b_pendings) (<[i:=(l, Done, false)]> slots)) deqs) pref = + map (get_value slots deqs) pref. +Proof. + intros Hpref HND. + induction pref as [|pref_hd pref IH]; first done. + assert (NoDup (pref ++ i :: b_pendings)) as HND_IH. + { simpl in HND. apply NoDup_cons in HND as [_ HND]. done. } + assert (∀ k, k ∈ pref → was_committed <$> slots !! k = Some true ∧ + k ∉ deqs) as Hpref_IH. + { intros k Hk. by apply Hpref, elem_of_list_further, Hk. } + rewrite /= IH; try done. clear IH HND_IH Hpref_IH. f_equal. + assert (i ≠ pref_hd) as Hi_not_pref_hd. + { simpl in HND. apply NoDup_cons in HND as (HND & _). + apply not_elem_of_app in HND as (_ & HND). + by apply not_elem_of_cons in HND as (HND & _). } + rewrite /get_value map_lookup_imap lookup_insert_ne; last done. + destruct (slots !! pref_hd) as [[[lp sp] wp]|]; last done. + destruct sp; try done. rewrite /= /helped /=. + rewrite decide_False; first done. + simpl in HND. apply NoDup_cons in HND as (HND & _). + apply not_elem_of_app in HND as (_ & HND). + by apply not_elem_of_cons in HND as (_ & HND). +Qed. + +Lemma annoying_lemma_2 slots deqs pref i l b_pendings : + block_valid slots (i, b_pendings) → + NoDup (pref ++ i :: b_pendings) → + map (get_value (map_imap (helped b_pendings) (<[i:=(l, Done, false)]> slots)) deqs) b_pendings = + get_values (<[i:=(l, Done, false)]> slots) b_pendings. +Proof. + intros (Hvalid_1 & Hvalid_2) HND. + induction b_pendings as [|p ps IH]; first done. simpl in *. + assert (i ≠ p) as Hi_not_p. + { intros ->. apply NoDup_app in HND as (_ & _ & HND). + apply NoDup_cons in HND as (HND & _). by set_solver +HND. } + rewrite lookup_insert_ne; last done. + assert (p ∈ p :: ps) as Hcomm%Hvalid_2 by set_solver. + destruct (slots !! p) + as [[[lp sp] wp]|] eqn:Hslots_p; [ f_equal | by inversion Hcomm ]. + - rewrite /= map_imap_insert /helped /= /get_value. + rewrite lookup_insert_ne; last done. rewrite map_lookup_imap Hslots_p /=. + destruct sp; try done. rewrite decide_True; [ done | by set_solver ]. + - rewrite -IH; first last; try done. + { apply NoDup_app in HND as (HND1 & HND2 & HND3). + apply NoDup_app. split; first done. split. + - intros e He. apply HND2 in He. apply not_elem_of_cons. + split; by set_solver +He. + - apply NoDup_cons in HND3 as (HND3_1 & HND3_2). + apply NoDup_cons. split; first by set_solver +HND3_1. + apply NoDup_cons in HND3_2 as (HND3_2_1 & HND3_2_2). done. } + { intros k Hk. by apply Hvalid_2, elem_of_list_further, Hk. } + apply map_ext_in. intros k Hk. + rewrite /get_value map_lookup_imap map_lookup_imap. + assert (i ≠ k) as Hi_not_k. + { intros ->. apply NoDup_app in HND as (_ & _ & HND). + apply NoDup_cons in HND as (HND & _). + apply not_elem_of_cons in HND as (_ & HND). + by apply HND, elem_of_list_In, Hk. } + rewrite lookup_insert_ne; last done. + assert (k ∈ p :: ps) as Hk_p_ps + by by apply elem_of_list_further, elem_of_list_In. + specialize (Hvalid_2 _ Hk_p_ps) as Hcomm_k. + destruct (slots !! k) as [[[lk sk] wk]|]; last by inversion Hcomm_k. + destruct sk; try done. rewrite /= /helped /=. + rewrite decide_True; last done. + rewrite decide_True; [ done | by apply elem_of_list_In ]. +Qed. + +Lemma big_lemma γe γs (ls : list loc) slots (p : list nat) : + NoDup p → + (∀ i, i ∈ p → was_committed <$> slots !! i = Some false) → + (own γs (● (of_slot_data <$> slots) : slotUR) -∗ + ([∗ map] i ↦ d ∈ slots, per_slot_own γe γs i d) -∗ + own γe (● (Excl' ls)) ={⊤ ∖ ↑N}=∗ + own γs (● (of_slot_data <$> map_imap (helped p) slots) : slotUR) ∗ + ([∗ map] i ↦ d ∈ map_imap (helped p) slots, per_slot_own γe γs i d) ∗ + own γe (● (Excl' (ls ++ get_values slots p))))%I. +Proof. + revert p. iIntros (p). + iInduction p as [|n ps] "IH" forall (slots ls); iIntros (HNoDup H) "Hs● Hbig He●". + - iModIntro. rewrite /= -app_nil_end map_imap_helped_nil. iFrame. + - assert (∀ i : nat, i ∈ ps → was_committed <$> slots !! i = Some false) as H1. + { intros i Hi. apply H. apply elem_of_list_further, Hi. } + assert (was_committed <$> slots !! n = Some false) as H2. + { apply H. apply elem_of_list_here. } + assert (∃ ln γn wn, slots !! n = Some (ln, Pend γn, wn)) as Hn. + { destruct (slots !! n) as [[[ln sn] wn]|]; last by inversion H2. + (destruct sn as [γn|γn|]; last by inversion H2); by exists ln, γn, wn. } + apply NoDup_cons in HNoDup. destruct HNoDup as [Hn_not_in_ps HNoDup]. + destruct Hn as [l [γ [w Hn]]]. + assert (slots = <[n:=(l, Pend γ, w)]> (delete n slots)) as Hs. + { by rewrite insert_delete insert_id. } + rewrite [in ([∗ map] _ ↦ _ ∈ slots, _)%I]Hs. + iDestruct (big_sepM_insert with "Hbig") + as "[Hbig_n Hbig]"; first by apply lookup_delete. + iDestruct "Hbig_n" as "[Hval_wit_n [Hwritten_n [Hpending_tok_n H]]]". + iDestruct "H" as (Q) "[Hsaved AU]". + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ (ls ++ [l]) with "He● He◯") as "[He● He◯]". + iMod ("Hclose" with "[$He◯]") as "HPost". + iMod (use_pending_tok with "Hs● Hpending_tok_n") + as "[Hs● Hcommitted_wit_n]"; first by rewrite Hn. + iCombine "Hsaved HPost" as "Hn". + iDestruct (big_sepM_insert _ (delete n slots) n (l, Help γ, w) + with "[Hn Hval_wit_n Hwritten_n Hcommitted_wit_n Hbig]") + as "Hbig"; first by apply lookup_delete. + { iClear "IH". iFrame "Hbig". rewrite /per_slot_own /=. iFrame. + iExists Q. iDestruct "Hn" as "[$ HPost]". iNext. done. } + rewrite insert_delete /update_slot Hn insert_delete. + assert (∀ i : nat, i ∈ ps → was_committed <$> <[n:=(l, Help γ, w)]> slots !! i = Some false) as HHH. + { intros i Hi. rewrite lookup_insert_ne; [ by apply H1 | by set_solver ]. } + iMod ("IH" $! (<[n:=(l, Help γ, w)]> slots) (ls ++ [l]) HNoDup HHH + with "Hs● Hbig He●") as "[Hs● [Hbig He●]]"; iClear "IH". + assert (map_imap (helped ps) (<[n:=(l, Help γ, w)]> slots) + = map_imap (helped (n :: ps)) slots) as ->. + { apply map_eq. intros i. destruct (decide (i = n)) as [->|Hi_not_n]. + - rewrite map_lookup_imap map_lookup_imap /= lookup_insert Hn /=. + rewrite /helped /=. rewrite decide_True; first done. set_solver. + - rewrite map_lookup_imap map_lookup_imap /= lookup_insert_ne; last done. + destruct (slots !! i) as [[[li si] wi]|]; last done. simpl. + rewrite /helped /=. destruct si; try done. + destruct (decide (i ∈ n :: ps)). + + rewrite decide_True; first done. set_solver. + + rewrite decide_False; first done. set_solver. } + iModIntro. iFrame. + by rewrite /= Hn -app_assoc /= get_values_not_in. +Qed. + +Lemma array_contents_cases γs slots deqs i li : + own γs (● (of_slot_data <$> slots) : slotUR) -∗ + slot_val_wit γs i li -∗ + ⌜array_get slots deqs i = SOMEV #li ∨ array_get slots deqs i = NONEV⌝. +Proof. + iIntros "Hs● Hval_wit_i". + iDestruct (use_val_wit with "Hs● Hval_wit_i") as %Hslots_i. + destruct (slots !! i) as [d|] eqn:HEq; last by inversion Hslots_i. + destruct d as [[li' si] wi]. inversion Hslots_i as [H]; subst li'. + rewrite /array_get HEq. simpl. iPureIntro. + destruct (decide (i ∈ deqs)); first by right. + destruct wi; by [ left | right ]. +Qed. + +(** * Specification of the queue operations *********************************) + +Lemma new_queue_spec sz : + 0 < sz → + {{{ True }}} + new_queue #sz + {{{ v γ, RET v; is_hwq sz γ v ∗ hwq_cont γ [] }}}. +Proof. + iIntros (Hsz Φ) "_ HΦ". wp_lam. wp_apply wp_fupd. + (** Allocate [q.ar], [q.back] and [q.p]. *) + wp_apply wp_allocN; [ lia | done | iIntros (ℓa) "[Hℓa _]" ]. + wp_alloc ℓb as "Hℓb". + wp_apply wp_new_proph; [ done | iIntros (rs p) "Hp" ]. + wp_pures. + (* Allocate the remaining ghost state. *) + iMod new_back as (γb) "Hb●". + iMod new_back as (γi) "Hi●". (* FIXME not about back. *) + iMod (new_elts []) as (γe) "[He● He◯]". + iMod new_no_contra as (γc) "HC". + iMod new_slots as (γs) "Hs●". + (* Allocate the invariant. *) + iMod (inv_alloc N _ (inv_hwq sz γb γi γe γc γs ℓa ℓb p) + with "[Hℓa Hℓb Hp Hb● Hi● He● HC Hs●]") as "#InvN". + { pose (pvs := proph_data sz ∅ rs). + pose (cont := NoCont (map (λ i, (i, [])) pvs)). + iNext. iExists 0%nat, pvs, [], [], cont, ∅, ∅. + rewrite array_content_empty Nat2Z.id fmap_empty /=. + iFrame. iSplitL. { iExists rs. by iFrame. } + repeat (iSplit; first done). iPureIntro. + repeat split_and; try done. + - intros i. split; intros Hi; [ by lia | by inversion Hi]. + - intros e He. set_solver. + - apply proph_data_NoDup. + - intros i. apply proph_data_sz. + - intros b. apply initial_block_valid. + - simpl. apply flatten_blocks_initial. } + (* Wrap things up. *) + iModIntro. iApply "HΦ". iFrame. + iExists γb, γi, γc, γs, ℓa, ℓb, p. by iSplit. +Qed. + +Lemma enqueue_spec sz γe (q : val) (l : loc) : + is_hwq sz γe q -∗ + <<< ∀ (ls : list loc), hwq_cont γe ls >>> + enqueue q #l @ ⊤ ∖ ↑N + <<< hwq_cont γe (ls ++ [l]), RET #() >>>. +Proof. + iIntros "Hq" (Φ) "AU". + iDestruct "Hq" as (γb γi γc γs ℓ_ar ℓ_back p ->) "#Inv". + rewrite /enqueue. wp_pures. wp_bind (FAA _ _)%E. + (* Open the invariant to perform the increment. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [Hb● [Hi● [He● [Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + wp_faa. assert (back + 1 = S back) as -> by lia. + iMod (back_incr with "Hb●") as "Hb●". + iAssert (i2_lower_bound γi match cont with + | WithCont _ i2 => i2 + | NoCont _ => (back `min` sz)%nat + end -∗ |==> + i2_lower_bound γi match cont with + | WithCont _ i2 => i2 + | NoCont _ => ((S back) `min` sz)%nat + end)%I as "Hup". + { destruct cont as [i1 i2|bs]; iIntros "Hi●"; first done. + iMod (i2_lower_bound_update with "Hi●") as "$"; [ lia | done ]. } + iMod ("Hup" with "Hi●") as "Hi●". + (* We first handle the case where there is no more space in the queue. *) + destruct (decide (back < sz)) as [Hback_sz|Hback_sz]; last first. + { iModIntro. iClear "AU". iSplitL. + - iNext. iExists (S back), pvs, pref, rest, cont, slots, deqs. + assert (S back `min` sz = back `min` sz)%nat as -> by lia. + iFrame. iPureIntro. repeat split_and; try done. + destruct cont as [i1 i2|bs]; last done. + destruct Hcont as ((H1 & H2) & H3 & H4). + by repeat (split; first lia). + - wp_pures. rewrite (bool_decide_false _ Hback_sz). wp_apply wp_diverge. } + (* We now have a reserved slot [i], which is still free. *) + pose (i := back). pose (elts := map (get_value slots deqs) pref ++ rest). + assert (slots !! back = None) as Hi_free. + { destruct (Hslots i) as [H1 H2]. rewrite min_l in H1; last by lia. + assert (¬ is_Some (slots !! back)). { intro H. apply H2 in H. lia. } + apply eq_None_not_Some, H. } + (* Useful fact: our index was not yet dequeued. *) + assert (i ∉ deqs) as Hi_not_in_deq. + { intros H. apply Hdeqs in H as (H & _). rewrite Hi_free in H. inversion H. } + (* We then handle the case where there is a contradiction going on. *) + destruct cont as [i1 i2|bs]. + { (* We access the atomic update and commit the element. *) + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ (elts ++ [l]) with "He● He◯") as "[He● He◯]". + iMod ("Hclose" with "[$He◯]") as "HΦ". + (* We allocate the new slot. *) + iMod (alloc_done_slot γs slots i l Hi_free with "Hs●") + as "[Hs [Htok_i [#val_wit_i [#commit_wit_i Hwriting_tok_i]]]]". + (* We also remember that we had contradiciton states. *) + iDestruct "Hcont" as "#cont_wit". + (* And we can close the invariant. *) + iModIntro. iSplitR "HΦ Hwriting_tok_i". + { iNext. iExists (S back), pvs, pref, (rest ++ [l]), (WithCont i1 i2). + iExists (<[i := (l, Done, false)]> slots), deqs. + rewrite fmap_insert /= array_content_NONEV; try done. iFrame. + iFrame. iSplitL "He●". + { rewrite /elts app_assoc map_get_value_not_in_pref; try done. + intros Hi%Hpref. rewrite Hi_free in Hi. destruct Hi; done. } + iSplitL "Hbig Htok_i". + { iApply big_sepM_insert. + + apply eq_None_not_Some. intros H. apply Hslots in H. lia. + + iFrame "Hbig". repeat (iSplit; first done). done. } + iFrame "cont_wit". + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + iPureIntro. repeat split_and; try done; try by lia. + - intros k. destruct sz; first by lia. + split; intros Hk. + + destruct (decide (k = i)) as [->|k_not_i]. + * rewrite lookup_insert. by eexists. + * rewrite lookup_insert_ne; last done. apply Hslots. by lia. + + destruct (decide (k = i)) as [->|k_not_i]. + * destruct sz; by lia. + * rewrite lookup_insert_ne in Hk; last done. + apply Hslots in Hk. by lia. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + by rewrite lookup_insert. + + rewrite lookup_insert_ne; last done. apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|HNeq]. + + split; first by rewrite lookup_insert. split; first done. + intros ->. apply Hpref in Hk as (_ & _ & H). done. + + rewrite lookup_insert_ne; last done. apply Hpref, Hk. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + by rewrite lookup_insert. + + rewrite /array_get. rewrite lookup_insert_ne; last done. + apply Hdeqs in Hk as (H1 & H2 & H3). repeat (split; first done). + rewrite /array_get in H3. + destruct (slots !! k) as [[[dl ds] dw]|]; last done. done. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + by rewrite lookup_insert. + + by rewrite lookup_insert_ne. + - rewrite /array_get lookup_insert_ne; first done. lia. + - rewrite /array_get lookup_insert_ne; last by lia. + destruct (slots !! i1) as [[[li1 si1] wi2]|]; last by inversion HC4. + rewrite decide_False; last done. inversion HC5; subst wi2. done. } + (* Let's clean up the context a bit. *) + clear Hslots Hstate Hpref Hdeqs Hcont Hi_not_in_deq Hi_free Hpvs_ND Hpvs_sz. + clear elts pvs pref rest slots deqs. subst i. rename back into i. + (* We can now move to the store. *) + wp_pures. rewrite (bool_decide_true _ Hback_sz). + wp_pures. wp_bind (_ <- _)%E. + (* We open the invariant again for the store. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [Hb● [Hi● [He● [>Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [>Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* Using witnesses, we show that our value and state have not changed. *) + iDestruct (use_val_wit with "Hs● val_wit_i") as %Hval_wit_i. + iDestruct (use_committed_wit with "Hs● commit_wit_i") as %Hval_commit_i. + iDestruct (writing_tok_not_written with "Hs● Hwriting_tok_i") as %Hnot_written_i. + (* We also show that the same contradiction ist still going on. *) + destruct cont as [i1' i2'|bs]; last first. + { by iDestruct (contra_not_no_contra with "Hcont cont_wit") as %Absurd. } + iDestruct (contra_agree with "cont_wit Hcont") as %[-> ->]. + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + (* Our slot is mapped. *) + assert (is_Some (slots !! i)) as Hslots_i. + { destruct (slots !! i) as [d|]; first by exists d. inversion Hval_wit_i. } + (* Our index is in the array. *) + assert (i < back `min` sz)%nat as Hi_le_back by by apply Hslots. + (* An we perform the store. *) + wp_apply (wp_store_offset _ _ ℓ_ar i (array_content sz slots deqs) with "Hℓ_ar"). + { apply array_content_is_Some. by lia. } + iIntros "Hℓ_ar". + (* We perform some updates. *) + iMod (use_writing_tok with "Hs● Hwriting_tok_i") as "[Hs● #written_wit_i]". + iModIntro. iSplitR "HΦ"; last by wp_pures. iNext. + (* It remains to re-establish the invariant. *) + pose (new_slots := update_slot i set_written slots). + iExists back, pvs, pref, rest, (WithCont i1 i2), new_slots, deqs. + subst new_slots. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written; + [ by iFrame | by lia | done | by apply Hstate ]. } + iSplitL "He●". + { erewrite map_ext; first by iFrame. rewrite /get_value. intros k. + destruct (decide (k = i)) as [->|Hk_not_i]. + - rewrite update_slot_lookup. destruct Hslots_i as [d Hslots_i]. + destruct d as [[ld sd] wd]. rewrite Hslots_i in Hnot_written_i. + inversion Hnot_written_i; subst wd. rewrite Hslots_i /=. done. + - rewrite update_slot_lookup_ne; last done. done. } + iSplitL "Hbig". + { rewrite /update_slot. destruct (slots !! i) as [d|] eqn:HEq; last done. + iApply big_sepM_insert; first by rewrite lookup_delete. + assert (slots = <[i:=d]> (delete i slots)) as HEq_slots. + { rewrite insert_delete. by rewrite insert_id. } + rewrite [X in ([∗ map] _ ↦ _ ∈ X, _)%I] HEq_slots. + iDestruct (big_sepM_insert with "Hbig") + as "[[H1 [H2 H3]] $]"; first by rewrite lookup_delete. + rewrite /per_slot_own val_of_set_written state_of_set_written. + iFrame. by rewrite was_written_set_written. } + iPureIntro. + destruct Hslots_i as [[[li si] wi] Hslots_i]. + repeat split_and; try done. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + rewrite update_slot_lookup. split; intros H; last done. + rewrite Hslots_i. by eexists. + + rewrite update_slot_lookup_ne; last done. by apply Hslots. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + rewrite update_slot_lookup Hslots_i /=. split; intros H. + * exfalso. rewrite Hslots_i in Hval_commit_i. + destruct si as [γ|γ|]; try by inversion Hval_commit_i. + * by inversion H. + + rewrite update_slot_lookup_ne; last done. apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hslots_i /=. repeat split. + * rewrite Hslots_i in Hval_commit_i. + destruct si; try by inversion Hval_commit_i. + * intros Hi%Hdeqs. destruct Hi as [H _]. + rewrite Hnot_written_i in H. inversion H. + * by apply Hpref in Hk as (_ & _ & H). + + rewrite update_slot_lookup_ne; last done. apply Hpref, Hk. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hslots_i /update_slot /=. + rewrite Hslots_i /= insert_delete /array_get lookup_insert. + rewrite decide_True; last done. repeat split; try done. + destruct si; try done. rewrite Hslots_i in Hval_commit_i. done. + + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + rewrite update_slot_lookup Hslots_i /=. + rewrite Hslots_i in HC4. by inversion HC4. + + by rewrite update_slot_lookup_ne. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + rewrite /array_get update_slot_lookup Hslots_i /=. + destruct (decide (i ∈ deqs)) as [H|H]; last done. + exfalso. apply Hdeqs in H as (H1 & H2 & H3). + rewrite Hnot_written_i in H1. inversion H1. + + by rewrite /array_get update_slot_lookup_ne. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + rewrite /array_get update_slot_lookup Hslots_i /=. + rewrite Hslots_i in HC5. inversion HC5; subst wi. + by rewrite decide_False. + + rewrite /array_get update_slot_lookup_ne; last done. + destruct (slots !! i1) as [[[li1 si1] wi1]|]; last by inversion HC4. + rewrite decide_False; last done. inversion HC5; subst wi1. done. } + (* There is no [Contra1]/[Contra2], first assume the prophecy is trivial. *) + destruct bs as [|b blocks]. + { (* We access the atomic update and commit the element. *) + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ (elts ++ [l]) with "He● He◯") as "[He● He◯]". + iMod ("Hclose" with "[$He◯]") as "HΦ". + (* We allocate the new slot. *) + iMod (alloc_done_slot γs slots i l Hi_free with "Hs●") + as "[Hs [Htok_i [#val_wit_i [#commit_wit_i Hwriting_tok_i]]]]". + (* And we can close the invariant. *) + iModIntro. iSplitR "HΦ Hwriting_tok_i". + { iNext. iExists (S back), pvs, pref, (rest ++ [l]), (NoCont []). + iExists (<[i := (l, Done, false)]> slots), deqs. + rewrite array_content_NONEV. iFrame. + iFrame. iSplitL "He●". + { rewrite /elts app_assoc map_get_value_not_in_pref; try done. + intros Hi%Hpref. rewrite Hi_free in Hi. destruct Hi; done. } + iSplitL "Hbig Htok_i". + { iApply big_sepM_insert. + + apply eq_None_not_Some. intros H. apply Hslots in H. lia. + + iFrame "Hbig". repeat (iSplit; first done). done. } + destruct Hcont as (HC1 & HC2 & HC3). + iPureIntro. repeat split_and; try done; try by lia. + - intros k. destruct sz; first by lia. + split; intros Hk. + + destruct (decide (k = i)) as [->|k_not_i]. + * rewrite lookup_insert. by eexists. + * rewrite lookup_insert_ne; last done. apply Hslots. by lia. + + destruct (decide (k = i)) as [->|k_not_i]. + * destruct sz; by lia. + * rewrite lookup_insert_ne in Hk; last done. + apply Hslots in Hk. by lia. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + by rewrite lookup_insert. + + rewrite lookup_insert_ne; last done. apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + by rewrite lookup_insert. + + rewrite lookup_insert_ne; last done. apply Hpref, Hk. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + by rewrite lookup_insert. + + rewrite /array_get. rewrite lookup_insert_ne; last done. + apply Hdeqs in Hk as (H1 & H2 & H3). repeat (split; first done). + rewrite /array_get in H3. + destruct (slots !! k) as [[[dl ds] dw]|]; last done. done. + - intros b Hb. by inversion Hb. + - done. + - done. + - done. } + (* Let's clean up the context a bit. *) + clear Hslots Hstate Hpref Hdeqs Hcont Hi_not_in_deq Hi_free Hpvs_ND Hpvs_sz. + clear pvs pref rest slots deqs elts. subst i. rename back into i. + (* We can now move to the store. *) + wp_pures. rewrite (bool_decide_true _ Hback_sz). + wp_pures. wp_bind (_ <- _)%E. + (* We open the invariant again for the store. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [Hb● [Hi● [He● [>Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [>Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* Using witnesses, we show that our value and state have not changed. *) + iDestruct (use_val_wit with "Hs● val_wit_i") as %Hval_wit_i. + iDestruct (use_committed_wit with "Hs● commit_wit_i") as %Hval_commit_i. + iDestruct (writing_tok_not_written with "Hs● Hwriting_tok_i") as %Hnot_written_i. + (* Our slot is mapped. *) + assert (is_Some (slots !! i)) as Hslots_i. + { destruct (slots !! i) as [d|]; first by exists d. inversion Hval_wit_i. } + (* Our index is in the array. *) + assert (i < back `min` sz)%nat as Hi_le_back by by apply Hslots. + (* An we perform the store. *) + wp_apply (wp_store_offset _ _ ℓ_ar i (array_content sz slots deqs) with "Hℓ_ar"). + { apply array_content_is_Some. by lia. } + iIntros "Hℓ_ar". + (* We perform some updates. *) + iMod (use_writing_tok with "Hs● Hwriting_tok_i") as "[Hs● #written_wit_i]". + iModIntro. iSplitR "HΦ"; last by wp_pures. iNext. + (* It remains to re-establish the invariant. *) + pose (new_slots := update_slot i set_written slots). + iExists back, pvs, pref, rest, cont, new_slots, deqs. + subst new_slots. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written; + [ by iFrame | by lia | done | by apply Hstate ]. } + iSplitL "He●". + { erewrite map_ext; first by iFrame. rewrite /get_value. intros k. + destruct (decide (k = i)) as [->|Hk_not_i]. + - rewrite update_slot_lookup. destruct Hslots_i as [d Hslots_i]. + destruct d as [[ld sd] wd]. rewrite Hslots_i in Hnot_written_i. + inversion Hnot_written_i; subst wd. rewrite Hslots_i /=. done. + - rewrite update_slot_lookup_ne; last done. done. } + iSplitL "Hbig". + { rewrite /update_slot. destruct (slots !! i) as [d|] eqn:HEq; last done. + iApply big_sepM_insert; first by rewrite lookup_delete. + assert (slots = <[i:=d]> (delete i slots)) as HEq_slots. + { rewrite insert_delete. by rewrite insert_id. } + rewrite [X in ([∗ map] _ ↦ _ ∈ X, _)%I] HEq_slots. + iDestruct (big_sepM_insert with "Hbig") + as "[[H1 [H2 H3]] $]"; first by rewrite lookup_delete. + rewrite /per_slot_own val_of_set_written state_of_set_written. + iFrame. by rewrite was_written_set_written. } + iPureIntro. + destruct Hslots_i as [[[li si] wi] Hslots_i]. + repeat split_and; try done. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + rewrite update_slot_lookup. split; intros H; last done. + rewrite Hslots_i. by eexists. + + rewrite update_slot_lookup_ne; last done. by apply Hslots. + - intros k. destruct (decide (k = i)) as [->|k_not_i]. + + rewrite update_slot_lookup Hslots_i /=. split; intros H. + * exfalso. rewrite Hslots_i in Hval_commit_i. + destruct si as [γ|γ|]; try by inversion Hval_commit_i. + * by inversion H. + + rewrite update_slot_lookup_ne; last done. apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hslots_i /=. repeat split. + * rewrite Hslots_i in Hval_commit_i. + destruct si; try by inversion Hval_commit_i. + * by intros Hi%Hpref. + * by apply Hpref in Hk as (_ & _ & H). + + rewrite update_slot_lookup_ne; last done. apply Hpref, Hk. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hslots_i /update_slot /=. + rewrite Hslots_i /= insert_delete /array_get lookup_insert. + rewrite decide_True; last done. repeat split; try done. + destruct si; try done. rewrite Hslots_i in Hval_commit_i. done. + + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - destruct cont as [i1 i2|bs]. + + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & HC6). split; first done. + destruct (decide (i1 = i)) as [->|Hi1_not_i]. + * rewrite /array_get update_slot_lookup Hslots_i /=. + repeat split_and; try done. + ** rewrite Hslots_i in Hval_commit_i. destruct si; try done. + ** rewrite decide_False; first done. apply Hstate. done. + * rewrite /array_get update_slot_lookup_ne; last done. + rewrite /array_get in HC3. done. + + destruct Hcont as (HC1 & HC2 & HC3). repeat split_and; try done. + intros b Hb. apply HC1 in Hb as (Hb1 & Hb2). split. + * destruct (decide (b.1 = i)) as [Hb1_is_i|Hb1_not_i]. + ** rewrite -Hb1_is_i in Hslots_i. by rewrite Hslots_i in Hb1. + ** rewrite /update_slot Hslots_i insert_delete. + by rewrite lookup_insert_ne. + * intros k Hk. destruct (decide (k = i)) as [Hk_is_i|Hk_not_i]. + ** rewrite /update_slot Hslots_i insert_delete. subst k. + rewrite lookup_insert /=. rewrite Hslots_i in Hval_commit_i. + destruct (was_committed (li, si, true)) eqn:H; last done. + exfalso. apply Hb2 in Hk. rewrite Hslots_i in Hk. inversion Hk. + destruct si; try done. + ** rewrite /update_slot Hslots_i insert_delete. + rewrite lookup_insert_ne; last done. apply Hb2, Hk. } + (* There is no [Contra1]/[Contra2], and the prophecy is non-trivial. *) + destruct Hcont as (Hblocks & Hrest & Hpvs). + assert (rest = []) as -> by by apply Hrest. + rewrite -app_nil_end in elts. rewrite -app_nil_end. + destruct b as [b_unused b_pendings]. + (* We compare our index with the unused element of the prophecy. *) + destruct (decide (b_unused = i)) as [->|b_unused_not_i]. + + (* We are the non-committed element of the prophecy: commit the block. *) + (* We allocate the new slot. *) + iMod (alloc_done_slot γs slots i l Hi_free with "Hs●") + as "[Hs● [Htok_i [#val_wit_i [#commit_wit_i Hwriting_tok_i]]]]". + (* We then commit at our index. *) + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ (elts ++ [l]) with "He● He◯") as "[He● He◯]". + iMod ("Hclose" with "[$He◯]") as "HΦ". + (* Our prophecy block must be valid. *) + assert (block_valid slots (i, b_pendings)) + as Hb_valid by apply Hblocks, elem_of_list_here. + rewrite /block_valid /= in Hb_valid. + destruct Hb_valid as [Hb_valid1 Hb_valid2]. + (* We also need to commit for all indices in in [p_pendings] *) + assert (NoDup (i :: b_pendings)) as Hblock_ND. + { apply NoDup_app in Hpvs_ND as (H & _ & _). subst pvs. + apply NoDup_app in H as (_ & _ & H). simpl in H. + rewrite app_comm_cons in H. by apply NoDup_app in H as (H & _ & _). } + apply NoDup_cons in Hblock_ND as (Hi & HNoDup). + iAssert (per_slot_own γe γs i (l, Done, false)) with "[Htok_i]" as "Hi". + { rewrite /per_slot_own /=. eauto with iFrame. } + iDestruct (big_sepM_insert (per_slot_own γe γs) slots i (l, Done, false) + with "[Hi Hbig]") as "Hbig"; [ done | by iFrame | .. ]. + iMod (big_lemma _ _ _ _ b_pendings HNoDup with "Hs● Hbig He●") as "[Hs● [Hbig He●]]". + { intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + exfalso. apply Hi, Hk. + + rewrite lookup_insert_ne; last done. apply Hb_valid2, Hk. } + (* And then we can close the invariant. *) + iModIntro. iSplitR "HΦ Hwriting_tok_i". + { pose (new_pref := pref ++ i :: b_pendings). + pose (new_slots := map_imap (helped b_pendings) (<[i:=(l, Done, false)]> slots)). + iNext. iExists (S back), pvs, new_pref, [], (NoCont blocks), new_slots, deqs. + iFrame. iSplitL "Hℓ_ar". + { assert (array_content sz slots deqs = array_content sz new_slots deqs) as ->; last done. + apply array_content_ext. intros k Hk. rewrite /new_slots /array_get. + rewrite map_lookup_imap. destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite lookup_insert Hb_valid1 /helped /= decide_False. + - rewrite lookup_insert_ne; last done. + destruct (slots !! k) as [[[dl ds] dw]|]; last done. + rewrite /helped /=. destruct ds as [dγ|dγ|]. + + destruct dw; try done; by destruct (decide (k ∈ b_pendings)). + + by destruct dw. + + by destruct dw. } + iSplitL "He●". + { rewrite -app_nil_end /new_pref /elts map_app map_cons. + rewrite [in get_value new_slots deqs i]/get_value. + rewrite [in new_slots !! i]/new_slots. + rewrite map_lookup_imap lookup_insert /= -app_assoc cons_middle. + assert (NoDup (pref ++ i :: b_pendings)) as HND. + { apply NoDup_app in Hpvs_ND as (HND & _ & _). + rewrite cons_middle app_assoc. + rewrite Hpvs /= in HND. rewrite cons_middle in HND. + rewrite app_assoc app_assoc in HND. + by apply NoDup_app in HND as (HND & _ & _). } + rewrite annoying_lemma_1; try done. + assert (map (get_value new_slots deqs) b_pendings + = get_values (<[i:=(l, Done, false)]> slots) b_pendings) as ->. + - rewrite /new_slots. by eapply annoying_lemma_2. + - done. + - intros k Hk. by apply Hpref in Hk as (H1 & H2 & _). } + iPureIntro. repeat split_and; try done. + - intros k. rewrite /new_slots map_lookup_imap. split; intros Hk. + + destruct (decide (k = i)) as [->|Hk_not_i]. + * rewrite lookup_insert /helped /=. by eexists. + * rewrite lookup_insert_ne; last done. + assert (is_Some (slots !! k)) as [d ->] by (apply Hslots; lia). + by apply is_Some_helped. + + destruct (decide (k = i)) as [->|Hk_not_i]; first by lia. + rewrite lookup_insert_ne in Hk; last done. + assert (k < back `min` sz)%nat as H; last by lia. + apply Hslots. destruct (slots !! k) as [d|]; first by exists d. + by inversion Hk. + - intros k. rewrite /new_slots map_lookup_imap. + destruct (decide (k = i)) as [->|Hk_not_i]; + first by rewrite lookup_insert /helped /=. + rewrite lookup_insert_ne; last done. split; intros Hk. + + destruct (slots !! k) as [d|] eqn:HEq; last done. + assert (was_committed <$> Some d ≫= helped b_pendings k = was_committed <$> Some d) as HEq1. + { destruct d as [[dl []] dw]; simpl; simpl in Hk; by rewrite Hk. } + rewrite HEq1 -HEq in Hk. apply Hstate in Hk. rewrite HEq in Hk. + assert (was_written <$> Some d ≫= helped b_pendings k = was_written <$> Some d) as HEq2. + { destruct d as [[dl []] []]; simpl; simpl in Hk; try by inversion Hk. + rewrite /helped /=. destruct (decide (k ∈ b_pendings)); done. } + rewrite HEq2. by inversion Hk. + + destruct (slots !! k) as [d|] eqn:HEq; last done. + assert (was_written <$> Some d ≫= helped b_pendings k = was_written <$> Some d) as HEq1. + { by destruct d as [[dl []] dw]; rewrite /helped; destruct (decide (k ∈ b_pendings)). } + rewrite HEq1 -HEq in Hk. apply Hstate in Hk. done. + - intros k Hk. subst new_pref new_slots. apply elem_of_app in Hk as [Hk|Hk]. + { apply Hpref in Hk as (H1 & H2). split; last done. + rewrite map_imap_insert /=. destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite lookup_insert. + - rewrite lookup_insert_ne; last done. rewrite map_lookup_imap. + destruct (slots !! k) as [[[dl ds] dw]|]; last by inversion H1. + rewrite /= /helped. destruct ds as [dγ|dγ|]; try done. } + apply elem_of_cons in Hk as [Hk|Hk]. + { subst k. split; last done. by rewrite map_imap_insert /= lookup_insert. } + apply Hb_valid2 in Hk as Hb_valid2_k. split. + + rewrite map_lookup_imap. destruct (decide (k = i)) as [->|Hk_not_i]. + * by rewrite lookup_insert /=. + * rewrite lookup_insert_ne; last done. + destruct (slots !! k) as [[[kl ks] kw]|]; last by inversion Hb_valid2_k. + rewrite /= /helped. destruct ks; try done. by rewrite /= decide_True. + + apply Hstate in Hb_valid2_k. apply Hstate in Hb_valid2_k. done. + - intros k Hk. subst new_slots. rewrite /array_get map_lookup_imap. + assert (k ≠ i) as Hk_not_i. { intros ->. apply Hi_not_in_deq, Hk. } + rewrite lookup_insert_ne; last done. apply Hdeqs in Hk as (H1 & H2 & H3). + destruct (slots !! k) as [[[lk sk] wk]|] eqn:HEq; last by inversion H1. + inversion H1; subst wk. rewrite /=. repeat split_and; try by destruct sk. + destruct sk; try done; simpl. + + rewrite decide_True; first done. + rewrite /array_get HEq in H3. simpl in H3. + destruct (decide (k ∈ deqs)); first done. by inversion H3. + + rewrite decide_True; first done. + rewrite /array_get HEq in H3. simpl in H3. + destruct (decide (k ∈ deqs)); first done. by inversion H3. + - intros b Hk. subst new_slots. rewrite map_imap_insert /=. + assert (b ∈ (i, b_pendings) :: blocks) as H by set_solver +Hk. + assert (NoDup (i :: b_pendings ++ flatten_blocks blocks)) as HND. + { subst pvs. apply NoDup_app in Hpvs_ND as (HND & _ & _). + apply NoDup_app in HND as (_ & _ & HND). done. } + apply flatten_blocks_mem1 in Hk as Hk1. + apply Hblocks in H as (H1 & H2). split. + + assert (b.1 ≠ i) as Hb1_not_i. + { intros HEq. apply NoDup_cons in HND as [HND1 HND2]. apply HND1. + rewrite -HEq. apply elem_of_app. by right. } + rewrite lookup_insert_ne; last done. by rewrite map_lookup_imap H1. + + intros j Hj. assert (j ≠ i) as Hj_not_i. + { intros HEq. apply NoDup_cons in HND as [HND1 HND2]. apply HND1. + rewrite -HEq. apply elem_of_app. right. + apply (flatten_blocks_mem2 _ _ Hk _ Hj). } + rewrite lookup_insert_ne; last done. rewrite map_lookup_imap. + apply H2 in Hj as Hcomm. + destruct (slots !! j) as [[[lj sj] wj]|]; last by inversion Hj. + rewrite /= /helped. destruct sj; try done. simpl. + assert (j ∉ b_pendings); last by rewrite decide_False. + intros Hj_contra. apply NoDup_cons in HND as [_ HND]. + apply NoDup_app in HND. destruct HND as (HND1 & HND2 & HND3). + apply (HND2 _ Hj_contra). apply (flatten_blocks_mem2 _ _ Hk _ Hj). + - by rewrite Hpvs /= /new_pref app_comm_cons app_assoc. } + clear Hslots Hstate Hpref Hdeqs Hpvs Hrest Hblocks Hi_free Hi_not_in_deq. + clear Hpvs_ND Hpvs_sz Hb_valid1 Hb_valid2 HNoDup Hi elts pvs pref slots deqs. + clear blocks b_pendings. subst i. rename back into i. + wp_pures. rewrite bool_decide_true; last done. wp_pures. + wp_bind (_ <- _)%E. + (* We open the invariant again for the store. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [Hb● [Hi● [He● [>Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [>Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* Using witnesses, we show that our value and state have not changed. *) + iDestruct (use_val_wit with "Hs● val_wit_i") as %Hval_wit_i. + iDestruct (use_committed_wit with "Hs● commit_wit_i") as %Hval_commit_i. + iDestruct (writing_tok_not_written with "Hs● Hwriting_tok_i") as %Hnot_written_i. + (* Our slot is mapped. *) + assert (is_Some (slots !! i)) as Hslots_i. + { destruct (slots !! i) as [d|]; first by exists d. inversion Hval_wit_i. } + (* Our index is in the array. *) + assert (i < back `min` sz)%nat as Hi_le_back by by apply Hslots. + (* An we perform the store. *) + wp_apply (wp_store_offset _ _ ℓ_ar i (array_content sz slots deqs) with "Hℓ_ar"). + { apply array_content_is_Some. by lia. } + iIntros "Hℓ_ar". + (* We perform some updates. *) + iMod (use_writing_tok with "Hs● Hwriting_tok_i") as "[Hs● #written_wit_i]". + iModIntro. iSplitR "HΦ"; last by wp_pures. iNext. + (* It remains to re-establish the invariant. *) + { pose (new_slots := update_slot i set_written slots). + iExists back, pvs, pref, rest, cont, new_slots, deqs. + subst new_slots. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written; + [ by iFrame | by lia | done | by apply Hstate ]. } + iSplitL "He●". + { erewrite map_ext; first by iFrame. rewrite /get_value. intros k. + destruct (decide (k = i)) as [->|Hk_not_i]. + - rewrite update_slot_lookup. destruct Hslots_i as [d Hslots_i]. + destruct d as [[ld sd] wd]. rewrite Hslots_i in Hnot_written_i. + inversion Hnot_written_i; subst wd. rewrite Hslots_i /=. done. + - rewrite update_slot_lookup_ne; last done. done. } + iSplitL "Hbig". + { rewrite /update_slot. destruct (slots !! i) as [d|] eqn:HEq; last done. + iApply big_sepM_insert; first by rewrite lookup_delete. + assert (slots = <[i:=d]> (delete i slots)) as HEq_slots. + { rewrite insert_delete. by rewrite insert_id. } + rewrite [X in ([∗ map] _ ↦ _ ∈ X, _)%I] HEq_slots. + iDestruct (big_sepM_insert with "Hbig") + as "[[H1 [H2 H3]] $]"; first by rewrite lookup_delete. + rewrite /per_slot_own val_of_set_written state_of_set_written. + iFrame. by rewrite was_written_set_written. } + iPureIntro. + repeat split_and; try done. + - intros k. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup. split; intros Hk; last by lia. + by apply fmap_is_Some. + + rewrite update_slot_lookup_ne; last done. apply Hslots. + - intros k. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup. split; intros Hk; exfalso. + * destruct (slots !! i) as [[[li si] wi]|]; last by inversion Hk. + inversion_clear Hnot_written_i. destruct si; inversion Hk. + inversion Hval_commit_i. + * destruct (slots !! i) as [[[li si] wi]|]; by inversion Hk. + + rewrite update_slot_lookup_ne; last done. by apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup /=. split. + * destruct (slots !! i) as [[[li si] wi]|]; first done. + by inversion Hval_wit_i. + * apply Hpref, Hk. + + rewrite update_slot_lookup_ne; last done. by apply Hpref. + - intros k Hk. assert (k ≠ i) as Hk_not_i. + { intros ->. apply Hdeqs in Hk as (H1 & H2 & H3). + apply Hstate in Hnot_written_i. rewrite /array_get in H3. + destruct Hslots_i as [[[li si] wi] Hslots_i]. + rewrite Hslots_i decide_False in H3; last done. + rewrite Hslots_i in H1. inversion H1; subst wi. inversion H3. } + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - destruct cont as [i1 i2|bs]. + + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & HC6). + split; first done. repeat split_and; try done. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** rewrite update_slot_lookup. + destruct (slots !! i) as [[[li si] wi]|]; first done. + by inversion Hval_wit_i. + ** by rewrite update_slot_lookup_ne. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** rewrite /array_get update_slot_lookup. + destruct (slots !! i) as [[[li si] wi]|] eqn:HEq; try done. + ** by rewrite /array_get update_slot_lookup_ne. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** rewrite /array_get update_slot_lookup. + destruct (slots !! i) as [[[li si] wi]|] eqn:HEq; try done. + inversion HC3; subst wi. done. + ** rewrite /array_get update_slot_lookup_ne; last done. + destruct (slots !! i1) as [[[li1 si1] wi1]|] eqn:HEq; try done. + rewrite decide_False; last done. inversion HC3; subst wi1. done. + + destruct Hcont as (HC1 & HC2 & HC3). repeat split_and; try done. + destruct Hslots_i as [[[li si] wi] Hslots_i]. + intros b Hb. apply HC1 in Hb as (Hb1 & Hb2). split. + * destruct (decide (b.1 = i)) as [Hb1_is_i|Hb1_not_i]. + ** rewrite -Hb1_is_i in Hslots_i. rewrite Hb1 in Hslots_i. + by inversion Hslots_i. + ** by rewrite /update_slot Hslots_i insert_delete lookup_insert_ne. + * intros k Hk. destruct (decide (k = i)) as [Hk_is_i|Hk_not_i]. + ** rewrite /update_slot Hslots_i insert_delete. subst k. + rewrite lookup_insert /=. rewrite Hslots_i in Hval_commit_i. + destruct (was_committed (li, si, true)) eqn:H; last done. + exfalso. apply Hb2 in Hk. rewrite Hslots_i in Hk. inversion Hk. + destruct si; try done. + ** rewrite /update_slot Hslots_i insert_delete. + rewrite lookup_insert_ne; last done. apply Hb2, Hk. } + + (* We are not the first non-done element, we will give away our AU. *) + iMod (saved_prop_alloc (Φ #())) as (γs_i) "#Hγs_i". + iMod (alloc_pend_slot γs slots i l γs_i Hi_free with "Hs●") + as "[Hs● [Htok_i [#val_wit_i [Hpend_tok_i [Hname_tok_i Hwriting_tok_i]]]]]". + (* We close the invariant, storing our AU. *) + iModIntro. iSplitR "Htok_i Hname_tok_i Hwriting_tok_i". + { pose (new_bs := glue_blocks (b_unused, b_pendings) i blocks). + pose (new_slots := <[i:=(l, Pend γs_i, false)]> slots). + iNext. iExists (S back), pvs, pref, [], (NoCont new_bs), new_slots, deqs. + rewrite -app_nil_end. iFrame. iSplitL "Hℓ_ar". + { assert (array_content sz slots deqs = array_content sz new_slots deqs) as ->; last done. + apply array_content_ext. intros k Hk. rewrite /new_slots /array_get. + destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite Hi_free lookup_insert decide_False. + - rewrite lookup_insert_ne; last done. destruct (slots !! k) as [d|]; last done. + destruct d as [[dl ds] dw]. rewrite /helped /=. + destruct ds as [dγ|dγ|]; destruct dw; try done. } + iSplitL "He●". + { erewrite map_ext_in; first done. subst new_slots. + intros k Hk%elem_of_list_In. rewrite /get_value. + assert (k ≠ i); last by rewrite lookup_insert_ne. + intros ->. apply Hpref in Hk as (H1 & H2). + rewrite Hi_free in H1. inversion H1. } + iSplitL "Hbig Hpend_tok_i AU". + { iApply big_sepM_insert; first done. iFrame. iSplit; first done. + iExists (Φ #()). iFrame. done. } + iPureIntro. subst new_slots. repeat split_and; try done. + - intros k. destruct sz; first by lia. + split; intros Hk. + + destruct (decide (k = i)) as [->|k_not_i]. + * rewrite lookup_insert. by eexists. + * rewrite lookup_insert_ne; last done. apply Hslots. by lia. + + destruct (decide (k = i)) as [->|k_not_i]. + * destruct sz; by lia. + * rewrite lookup_insert_ne in Hk; last done. + apply Hslots in Hk. by lia. + - intros k. destruct (decide (k = i)) as [->|Hk_not_i]. + + by rewrite lookup_insert. + + rewrite lookup_insert_ne; last done. apply Hstate. + - intros k Hk. rewrite lookup_insert_ne; first by apply Hpref, Hk. + intros HEq. subst k. apply Hpref in Hk as [H _]. + rewrite Hi_free in H. inversion H. + - intros k Hk. rewrite /array_get lookup_insert_ne. + + apply Hdeqs in Hk. by rewrite /array_get in Hk. + + intros <-. apply Hdeqs in Hk as [Hk _]. rewrite Hi_free in Hk. done. + - intros b Hb. subst new_bs. rewrite Hpvs in Hpvs_ND. + apply NoDup_app in Hpvs_ND as (HND & _ & _). + apply NoDup_app in HND as (_ & _ & HND). simpl in HND. + by eapply glue_blocks_valid. + - subst pvs new_bs. f_equal. apply flatten_blocks_glue. } + clear Hslots Hstate Hpref Hdeqs Hblocks Hrest Hpvs Hi_free Hi_not_in_deq. + clear Hpvs_ND Hpvs_sz b_unused b_unused_not_i elts blocks pvs pref slots. + clear deqs b_pendings. subst i. rename back into i. + wp_pures. rewrite bool_decide_true; last done. + wp_pures. wp_bind (_ <- _)%E. + (* We open the invariant again for the store. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [Hb● [Hi● [He● [>Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [>Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* Using witnesses, we show that our value and state have not changed. *) + iDestruct (use_val_wit with "Hs● val_wit_i") as %Hval_wit_i. + iDestruct (writing_tok_not_written with "Hs● Hwriting_tok_i") as %Hnot_written_i. + (* Our slot is mapped. *) + assert (is_Some (slots !! i)) as Hslots_i. + { destruct (slots !! i) as [d|]; first by exists d. inversion Hval_wit_i. } + (* Our index is in the array. *) + assert (i < back `min` sz)%nat as Hi_le_back by by apply Hslots. + (* An we perform the store. *) + wp_apply (wp_store_offset _ _ ℓ_ar i (array_content sz slots deqs) with "Hℓ_ar"). + { apply array_content_is_Some. by lia. } + iIntros "Hℓ_ar". + (* We now look at the state of our cell. *) + destruct Hslots_i as [[[l' s] w] Hi]. + rewrite Hi in Hval_wit_i. simpl in Hval_wit_i. + inversion Hval_wit_i; subst l'. + destruct s as [γs_i'|γs_i'|]. + - (* We are still in the pending state: contradiction. *) + (* We need to run our atomic update ourselves, we recover it. *) + rewrite -[in X in ([∗ map] _ ↦ _ ∈ X, _)%I](insert_id _ _ _ Hi). + rewrite -insert_delete. + iDestruct (big_sepM_insert with "Hbig") + as "[Hbig_i Hbig]"; first by apply lookup_delete. + iDestruct "Hbig_i" as "[_ [_ [Hcommit_tok_i HAU]]]". + iDestruct "HAU" as (Q) "[Hsaved AU]". + (* We use the name token to show that γs_i and γs_i' are equal. *) + iDestruct (use_name_tok with "Hs● Hname_tok_i") as %Hname_tok_i. + assert (γs_i' = γs_i) as Hγs_i; last subst γs_i'. + { rewrite Hi /= in Hname_tok_i. by inversion Hname_tok_i. } + iDestruct (saved_prop_agree with "Hγs_i Hsaved") as "HQ_is_Φ". + (* We run our atomic update ourself. *) + pose (elts := map (get_value slots deqs) pref ++ rest). + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ (elts ++ [l]) with "He● He◯") as "[He● He◯]". + iMod ("Hclose" with "[$He◯]") as "HΦ". + iMod (use_writing_tok with "Hs● Hwriting_tok_i") as "[Hs● #written_wit_i]". + iMod (use_pending_tok with "Hs● Hcommit_tok_i") as "[Hs● #commit_wit_i]". + { by rewrite update_slot_lookup Hi /=. } + iMod (helped_to_done with "Hs● Hname_tok_i") as "Hs●". + { by rewrite update_slot_lookup update_slot_lookup Hi. } + (* We now act according ot the contradiction status. *) + destruct cont as [i1 i2|bs]. + * (* A contradiction has arised from somewhere else, we keep it. *) + iModIntro. iSplitR "HQ_is_Φ HΦ". + { iNext. iExists back, pvs, pref, (rest ++ [l]), (WithCont i1 i2). + iExists (update_slot i set_written_and_done slots), deqs. + subst elts. rewrite app_assoc. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written_and_done; + [ by iFrame | by lia | by rewrite Hi | by apply Hstate ]. } + iSplitL "He●". + { erewrite map_ext_in; first done. intros k Hk%elem_of_list_In. + rewrite /get_value /update_slot Hi insert_delete. + destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite lookup_insert Hi. + - by rewrite lookup_insert_ne. } + iSplitL "Hs●". + { repeat rewrite update_slot_update_slot. by rewrite /update_slot Hi. } + iSplitL. + { rewrite /update_slot Hi. + iApply big_sepM_insert; first by rewrite lookup_delete. + iFrame "Hbig". rewrite /per_slot_own /=. iFrame. + iSplit; first done. iSplit; done. } + iPureIntro. + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + repeat split_and; try lia; try done. + - intros k. destruct (decide (i = k)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hi. split; [ by eexists | lia ]. + + rewrite update_slot_lookup_ne; last done. apply Hslots. + - intros k. split; intros Hk. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne; last done. + rewrite update_slot_lookup_ne in Hk; last done. + by apply Hstate. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne in Hk; last done. by apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hi /=. split; [ done | by apply Hpref, Hk ]. + + rewrite update_slot_lookup_ne; last done. apply Hpref, Hk. + - intros k Hk. assert (k ≠ i) as Hk_not_i. + { intros ->. apply Hdeqs in Hk as (H1 & H2 & H3). + apply Hstate in Hnot_written_i. rewrite /array_get in H3. + rewrite Hi decide_False in H3; last done. + rewrite Hi in H1. inversion H1; subst w. inversion H3. } + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + by rewrite update_slot_lookup Hi. + + by rewrite update_slot_lookup_ne. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + by rewrite update_slot_lookup Hi /=. + + rewrite update_slot_lookup_ne; last done. + destruct (slots !! i1) as [[[li1 si1] wi1]|]; last by inversion HC4. + inversion HC5; subst wi1. done. + - destruct (decide (i1 = i)) as [->|Hi1_not_i]. + + by rewrite /array_get update_slot_lookup Hi /= decide_False. + + rewrite /array_get update_slot_lookup_ne; last done. + destruct (slots !! i1) as [[[li1 si1] wi1]|]; last by inversion HC4. + rewrite decide_False; last done. inversion HC5; subst wi1. done. } + wp_pures. iRewrite "HQ_is_Φ". done. + * (* No contradiction yet, make it ours if the prophecy is non-trivial. *) + iAssert (match bs with + | [] => i2_lower_bound γi (back `min` sz)%nat + | _ => no_contra γc ∗ i2_lower_bound γi (back `min` sz)%nat + end -∗ |==> + match bs with + | [] => True + | (i2, _) :: _ => contra γc i i2 + end ∗ + match bs with + | [] => i2_lower_bound γi (back `min` sz)%nat + | (i2, _) :: _ => i2_lower_bound γi i2 + end)%I as "Hup". + { destruct bs as [|[i2 ps] bs]; first (iIntros "Hi●"; by iFrame). + iIntros "[Hcont Hi●]". iMod (to_contra i i2 with "Hcont") as "$". + iMod (i2_lower_bound_update _ _ i2 with "Hi●") as "$"; last done. + assert (block_valid slots (i2, ps)) as [Hvalid _]. + { destruct Hcont as (Hblocks & _ & _). apply Hblocks, elem_of_list_here. } + assert (¬ (i2 < back `min` sz)%nat) as H%not_lt; last by lia. + eapply iffRLn. apply Hslots. intros H. rewrite Hvalid in H. by inversion H. } + iAssert (match bs with + | [] => i2_lower_bound γi (back `min` sz)%nat + | _ => no_contra γc ∗ i2_lower_bound γi (back `min` sz)%nat + end ∗ + match bs with + | [] => no_contra γc + | _ => True + end)%I with "[Hcont Hi●]" as "[HNC_triv HNC_non_triv]". + { destruct bs; by iFrame. } + iMod ("Hup" with "HNC_triv") as "[#HC_triv Hi●]". + (* We can now close the invariant. *) + iModIntro. iSplitR "HQ_is_Φ HΦ". + { pose (new_slots := update_slot i set_written_and_done slots). + pose (cont := match bs with [] => NoCont [] | (i2, _) :: _ => WithCont i i2 end). + iNext. iExists back, pvs, pref, (rest ++ [l]), cont, new_slots, deqs. + subst new_slots elts cont. rewrite app_assoc. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written_and_done; + [ by iFrame | by lia | by rewrite Hi | by apply Hstate ]. } + iSplitL "Hi●". + { destruct bs as [|[b_u b_ps] bs]; by iFrame. } + iSplitL "He●". + { erewrite map_ext_in; first done. intros k Hk%elem_of_list_In. + rewrite /get_value /update_slot Hi insert_delete. + destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite lookup_insert Hi. + - by rewrite lookup_insert_ne. } + iSplitL "Hs●". + { repeat rewrite update_slot_update_slot. by rewrite /update_slot Hi. } + iSplitR "HNC_non_triv". + { rewrite /update_slot Hi. + iApply big_sepM_insert; first by rewrite lookup_delete. + iFrame "Hbig". rewrite /per_slot_own /=. iFrame. + iSplit; first done. iSplit; done. } + iSplitL "HNC_non_triv"; first by destruct bs as [|[i2 ps] bs]. + iPureIntro. repeat split_and. + - intros k. destruct (decide (i = k)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hi. split; [ by eexists | lia ]. + + rewrite update_slot_lookup_ne; last done. apply Hslots. + - intros k. split; intros Hk. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne; last done. + rewrite update_slot_lookup_ne in Hk; last done. + by apply Hstate. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne in Hk; last done. by apply Hstate. + - intros k Hk. apply Hpref in Hk as (H1 & H2 & _). repeat split; try done. + + destruct (decide (k = i)) as [->|Hk_not_i]. + * by rewrite update_slot_lookup Hi. + * by rewrite update_slot_lookup_ne. + + destruct bs as [|[b_u b_ps] bs]; first done. + intros ->. rewrite Hi in H1. by inversion H1. + - intros k Hk. assert (k ≠ i) as Hk_not_i. + { intros ->. apply Hdeqs in Hk as (H1 & H2 & H3). + apply Hstate in Hnot_written_i. rewrite /array_get in H3. + rewrite Hi decide_False in H3; last done. + rewrite Hi in H1. inversion H1; subst w. inversion H3. } + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - done. + - done. + - destruct Hcont as (HC1 & HC2 & HC3). + destruct bs as [|[i2 ps] bs]. + + repeat split_and; try done. intros. by set_solver. + + repeat split_and; try lia. + * assert (i < back `min` sz)%nat + as Hi_lt by (apply Hslots; by eexists). + assert (block_valid slots (i2, ps)) + as Hvalid by apply HC1, elem_of_list_here. + assert (slots !! i2 = None) + as Hi2_None by by destruct Hvalid as (H & _). + assert (¬ i2 < back `min` sz)%nat as Hi2_ge; last by lia. + intros H%Hslots. rewrite Hi2_None in H. by inversion H. + * apply Hpvs_sz. subst pvs. apply elem_of_app. right. simpl. + by apply elem_of_list_here. + * by rewrite update_slot_lookup Hi /=. + * by rewrite update_slot_lookup Hi /=. + * by apply Hstate. + * rewrite /array_get update_slot_lookup Hi /=. + rewrite decide_False; first done. apply Hstate. done. + * rewrite HC3 /=. exists (ps ++ flatten_blocks bs). + by rewrite cons_middle app_assoc. } + wp_pures. iRewrite "HQ_is_Φ". done. + - (* We have moved to the helped state. *) + assert (slots = <[i := (l, Help γs_i', w)]> (delete i slots)) + as Hslots_i by by rewrite insert_delete insert_id. + rewrite [X in ([∗ map] _ ↦ _ ∈ X, _)%I]Hslots_i. + (* We recover our postcondition. *) + iDestruct (big_sepM_insert with "Hbig") + as "[Hbig_i Hbig]"; first by apply lookup_delete. + iDestruct "Hbig_i" as "[_ [_ [Hcommit_wit_i Hpost]]]". + iDestruct "Hpost" as (Q) "[Hsaved Hpost]". + (* We use the name token to show that γs_i and γs_i' are equal. *) + iDestruct (use_name_tok with "Hs● Hname_tok_i") as %Hname_tok_i. + assert (γs_i' = γs_i) as Hγs_i; last subst γs_i'. + { rewrite Hi /= in Hname_tok_i. by inversion Hname_tok_i. } + iDestruct (saved_prop_agree with "Hγs_i Hsaved") as "HQ_is_Φ". + (* We need to move from helped to done. *) + iMod (helped_to_done with "Hs● Hname_tok_i") as "Hs●". { by rewrite Hi. } + (* We perform some updates. *) + iMod (use_writing_tok with "Hs● Hwriting_tok_i") as "[Hs● #written_wit_i]". + iModIntro. iSplitR "HQ_is_Φ Hpost". + { pose (new_slots := update_slot i set_written_and_done slots). + iNext. iExists back, pvs, pref, rest, cont, new_slots, deqs. + subst new_slots. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_set_written_and_done; + [ by iFrame | by lia | by rewrite Hi | by apply Hstate ]. } + iSplitL "He●". + { erewrite map_ext_in; first done. intros k Hk%elem_of_list_In. + rewrite /get_value /update_slot Hi insert_delete. + destruct (decide (k = i)) as [->|Hk_not_i]. + - by rewrite lookup_insert Hi. + - by rewrite lookup_insert_ne. } + iSplitL "Hs●". + { repeat rewrite update_slot_update_slot. by rewrite /update_slot Hi. } + iSplitL. + { rewrite /update_slot Hi. + iApply big_sepM_insert; first by rewrite lookup_delete. + iFrame "Hbig". rewrite /per_slot_own /=. iFrame. iSplit; done. } + iPureIntro. repeat split_and; try done. + - intros k. destruct (decide (i = k)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hi. split; [ by eexists | lia ]. + + rewrite update_slot_lookup_ne; last done. apply Hslots. + - intros k. split; intros Hk. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne; last done. + rewrite update_slot_lookup_ne in Hk; last done. + by apply Hstate. + + assert (k ≠ i) as Hk_not_i. + { intros ->. by rewrite update_slot_lookup Hi in Hk. } + rewrite update_slot_lookup_ne in Hk; last done. by apply Hstate. + - intros k Hk. destruct (decide (k = i)) as [->|Hk_not_i]. + + rewrite update_slot_lookup Hi. split; first done. apply Hpref, Hk. + + rewrite update_slot_lookup_ne; last done. apply Hpref, Hk. + - intros k Hk. assert (k ≠ i) as Hk_not_i. + { intros ->. apply Hdeqs in Hk as (H1 & H2 & H3). + apply Hstate in Hnot_written_i. rewrite /array_get in H3. + rewrite Hi decide_False in H3; last done. + rewrite Hi in H1. inversion H1; subst w. inversion H3. } + rewrite /array_get update_slot_lookup_ne; last done. + apply Hdeqs in Hk. rewrite /array_get in Hk. done. + - destruct cont as [i1 i2|bs]. + + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & HC6). + split; first done. repeat split_and; try done. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** by rewrite update_slot_lookup Hi. + ** by rewrite update_slot_lookup_ne. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** by rewrite /array_get update_slot_lookup Hi /=. + ** rewrite /array_get update_slot_lookup_ne; last done. + rewrite /array_get in HC3. done. + * destruct (decide (i1 = i)) as [->|Hi1_not_i]. + ** by rewrite /array_get update_slot_lookup Hi decide_False. + ** rewrite /array_get update_slot_lookup_ne; last done. + rewrite /array_get in HC3. done. + + destruct Hcont as (HC1 & HC2 & HC3). repeat split_and; try done. + intros b Hb. apply HC1 in Hb as (H1 & H2). split. + ** assert (b.1 ≠ i) as Hb1_not_i. + { intros H. rewrite H in H1. by rewrite Hi in H1. } + by rewrite update_slot_lookup_ne. + ** intros k Hk. assert (k ≠ i) as Hb1_not_i. + { intros H. subst k. apply H2 in Hk. rewrite Hi in Hk. + by inversion Hk. } + rewrite update_slot_lookup_ne; last done. by apply H2. } + wp_pures. iRewrite "HQ_is_Φ". done. + - (* We are in the done state: contradiction. *) + iDestruct (big_sepM_lookup _ _ i with "Hbig") + as "[_ [_ H]]"; first done; simpl. + iDestruct "H" as "[_ Htok_i']". + by iDestruct (slot_token_exclusive with "Htok_i Htok_i'") as "H". +Qed. + +Lemma dequeue_spec sz γe (q : val) : + is_hwq sz γe q -∗ + <<< ∀ (ls : list loc), hwq_cont γe ls >>> + dequeue q @ ⊤ ∖ ↑N + <<< ∃ (l : loc) ls', ⌜ls = l :: ls'⌝ ∗ hwq_cont γe ls', RET #l >>>. +Proof. + iIntros "Hq" (Φ) "AU". iLöb as "IH". + iDestruct "Hq" as (γb γi γc γs ℓ_ar ℓ_back p ->) "#Inv". + wp_lam. wp_pures. wp_bind (! _)%E. + (* We need to open the invariant to read [q.back]. *) + iInv "Inv" as (back pvs pref rest cont slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [>Hb● [Hi● [He● [Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [>Hcont Hpures]]]". wp_load. + (* If there is a contradiction, remember that. *) + iAssert (match cont with + | NoCont _ => True + | WithCont i1 i2 => contra γc i1 i2 + end)%I with "[Hcont]" as "#Hinit_cont". + { destruct cont as [i1 i2|bs]; [ by iDestruct "Hcont" as "#C" | done ]. } + (* We remember the current back value. *) + iMod (back_snapshot with "Hb●") as "[Hb● Hback_snap]". + iMod (i2_lower_bound_snapshot with "Hi●") as "[Hi● Hi2_lower_bound]". + (* We close the invariant again. *) + iModIntro. iSplitR "AU Hback_snap Hi2_lower_bound". + { iNext. repeat iExists _. eauto with iFrame. } + clear pref rest slots deqs pvs. + (* The range is the min between [q.back - 1] and [q.size - 1]. *) + wp_bind (minimum _ _)%E. wp_apply minimum_spec_nat. wp_pures. + (* We now prove the inner loop part by induction in the index. *) + assert (back `min` sz ≤ back `min` sz)%nat as Hn by done. + assert (match cont with + | NoCont _ => True + | WithCont i1 _ => (back `min` sz - back `min` sz ≤ i1)%nat + end) as Hcont_i1 by (destruct cont as [i1 _|_]; lia). + revert Hn Hcont_i1. generalize (back `min` sz)%nat at 1 4 7 as n. + intros n Hn Hcont_i1. + iInduction n as [|n] "IH_loop" forall (Hn Hcont_i1). + (* The bas case is trivial. *) + { wp_lam. wp_pures. iApply "IH"; last done. + iExists _, _, _, _, _, _, _. iSplitR; done. } + (* Now the induction case: we need to open the invariant for the load. *) + wp_lam. wp_pures. wp_bind (! _)%E. + iInv "Inv" as (back' pvs pref rest cont' slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [>Hb● [Hi● [He● [Hs● HInv]]]]]]". + iDestruct "HInv" as "[Hproph [Hbig [Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* We use our snapshot to show that back is smaller that back'. *) + iDestruct (back_le with "Hb● Hback_snap") as %Hback. + (* We define the loop index as [i]. *) + pose (i := ((back `min` sz) - S n)%nat). + assert ((back `min` sz)%nat - S n = i) as -> by by rewrite Nat2Z.inj_sub. + (* We can now load. *) + wp_apply (wp_load_offset _ _ ℓ_ar i _ (array_get slots deqs i) with "Hℓ_ar"); + [ apply array_content_lookup; lia | iIntros "Hℓa" ]. + (* If there was an initial contradiction, it is still here. *) + iAssert ⌜match cont with + | NoCont _ => True + | WithCont i1 i2 => cont' = cont ∧ (back `min` sz - S n ≤ i1)%nat + end⌝%I as %Hinitial_cont. + { destruct cont as [i1 i2|bs]; destruct cont' as [i1' i2'|bs']; try done. + - iDestruct (contra_agree with "Hinit_cont Hcont") as %[-> ->]. + iPureIntro. split; first done. + destruct Hcont as (((H1 & H2) & H3) & _). lia. + - by iDestruct (contra_not_no_contra with "Hcont Hinit_cont") as "False". } + (* We then reason by cas on the physical contents of slot [i]. *) + destruct (decide (array_get slots deqs i = NONEV)) as [Hi_NULL|Hi_not_NULL]. + { rewrite Hi_NULL. iModIntro. + iSplitR "AU Hback_snap Hi2_lower_bound". + { iNext. repeat iExists _. iFrame. iSplit; done. } + wp_pures. assert (S n - 1 = n%nat) as -> by lia. + iApply ("IH_loop" with "[] [] AU Hback_snap"). + - iPureIntro. lia. + - iPureIntro. destruct cont as [i1 i2|bs]; last done. + destruct Hinitial_cont as [-> Hi1]. + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & HC6). + apply le_lt_or_eq in Hcont_i1 as [H|H]; rewrite -/i in H; first by lia. + exfalso. subst i1. + assert (is_Some (slots !! i)) as [d Hslots_i] by (apply Hslots; lia). + destruct d as [[li si] wi]. rewrite /array_get Hslots_i /= in Hi_NULL. + rewrite /array_get Hslots_i in HC3. rewrite decide_False in Hi_NULL; last done. + inversion HC3; subst wi. by inversion Hi_NULL. + - by iFrame. } + (* We know that a non-null value [li] at index [i], we get a witness. *) + assert (is_Some (slots !! i)) as [[[li si] wi] Hslots_i]. + { rewrite /array_get in Hi_not_NULL. + destruct (slots !! i) as [d|]; last done. by eexists. } + assert (array_get slots deqs i = SOMEV #li) as ->. + { rewrite /array_get Hslots_i /=. rewrite /array_get Hslots_i in Hi_not_NULL. + revert Hi_not_NULL. destruct (decide (i ∈ deqs)); intros H; first done. + by destruct wi. } + iMod (val_wit_from_auth γs i li with "Hs●") + as "[Hs● #Hval_wit_i]"; first by rewrite Hslots_i. + (* Close the invariant and clean up the context. *) + iModIntro. iSplitR "AU Hback_snap Hi2_lower_bound". + { iNext. repeat iExists _. iFrame. iSplit; done. } + clear Hslots Hstate Hpref Hdeqs Hcont Hinitial_cont Hback back' Hpvs_ND. + clear Hpvs_sz pvs pref rest cont' Hslots_i si wi Hi_not_NULL slots deqs. + (* Finally, the interesting where the cell was non-NULL on the load. *) + wp_pures. wp_bind (Resolve _ _ _)%E. + iInv "Inv" as (back' pvs pref rest cont' slots deqs) "HInv". + iDestruct "HInv" as "[Hℓ_back [Hℓ_ar [>Hb● [>Hi● [He● [>Hs● HInv]]]]]]". + iDestruct "HInv" as "[>Hproph [Hbig [>Hcont >Hpures]]]". + iDestruct "Hpures" as %(Hslots & Hstate & Hpref & Hdeqs & Hpvs_OK & Hcont). + destruct Hpvs_OK as (Hpvs_ND & Hpvs_sz). + (* If there was an initial contradiction, it is still here. *) + iAssert ⌜match cont with + | NoCont _ => True + | WithCont i1 i2 => cont' = cont ∧ (back `min` sz - S n ≤ i1)%nat + end⌝%I as %Hinitial_cont. + { destruct cont as [i1 i2|bs]; destruct cont' as [i1' i2'|bs']; try done. + - iDestruct (contra_agree with "Hinit_cont Hcont") as %[-> ->]. + iPureIntro. split; first done. destruct Hcont as (((H1 & H2) & H3) & _). done. + - by iDestruct (contra_not_no_contra with "Hcont Hinit_cont") as "False". } + (* We resolve. *) + iDestruct "Hproph" as (rs) "[Hp Hpvs]". iDestruct "Hpvs" as %Hpvs. + wp_apply (wp_resolve with "Hp"); first done. + (* We reason by case on the success of the CAS. *) + iDestruct (array_contents_cases γs slots deqs with "Hs● Hval_wit_i") as %[Hi|Hi]. + * (* The CmpXchg succeeded. *) iClear "IH_loop IH". + assert (array_content sz slots deqs !! i = Some (SOMEV #li)). + { rewrite array_content_lookup; last by lia. by rewrite Hi. } + wp_apply (wp_cmpxchg_suc_offset with "Hℓ_ar"); + [ done | done | by right | iIntros "Hℓ_ar" (rs' ->) "Hp" ]. + (* Note that [i] is used (otherwise the CmpXchg would have failed). *) + iDestruct (use_val_wit with "Hs● Hval_wit_i") as %Hval_wit_i. + iDestruct (back_le with "Hi● Hi2_lower_bound") as %Hi2. + assert (is_Some (slots !! i)) as [[[dl ds] dw] Hslots_i]. + { destruct (slots !! i) as [d|]; [ by exists d | by inversion Hval_wit_i ]. } + assert (dl = li) as Hdl_li; last subst dl. + { rewrite Hslots_i in Hval_wit_i. by inversion Hval_wit_i. } + (* We now reason by case on whether the enqueue at [i] was committed. *) + destruct (was_committed (li, ds, dw)) eqn:Hcommitted. + { (* We first consider the case where it was committed. *) + (* If [i] has been dequeued alread: contradiction. *) + assert (i ∉ deqs) as Hi_not_deq. + { intros Hi_deq. specialize (Hdeqs i Hi_deq) as (H1 & H2 & H3). + rewrite Hslots_i /= in H1. inversion H1; subst dw. + rewrite /array_get Hslots_i in Hi. rewrite decide_True in Hi; last done. + inversion Hi. } + (* We clean up the prophecy. *) + rewrite /= decide_True in Hpvs; last lia. rewrite Nat2Z.id in Hpvs. + rewrite decide_False in Hpvs; last done. + (* We then show that the commit prefix cannot be empty. *) + destruct pref as [|i' new_pref]. + { exfalso. destruct cont as [i1 i2|_]. + - destruct Hinitial_cont as [-> Hi1]. + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + rewrite Hpvs /= in HC8. destruct HC8 as [junk HEq]. + inversion HEq as [[HEq1 HEq2]]. lia. + - destruct cont' as [i1' i2'|bs]. + + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + rewrite Hpvs /= in HC8. destruct HC8 as [junk HEq]. + inversion HEq as [[HEq1 HEq2]]. + assert (back < i2')%nat; last by lia. lia. + + destruct Hcont as (HC1 & HC2 & HC3). rewrite Hpvs /= in HC3. + destruct bs as [|[b_u b_ps] bs]; first by inversion HC3. + simpl in HC3. inversion HC3 as [[HEq1 HEq2]]. + assert (block_valid slots (b_u, b_ps)) + as [Hvalid _] by apply HC1, elem_of_list_here. + rewrite /= -HEq1 Hslots_i in Hvalid. inversion Hvalid. } + assert (i' = i) as ->. + { destruct cont' as [i1' i2'|bs]. + - destruct Hcont as (_ & _ & _ & _ & _ & HC). + rewrite Hpvs in HC. destruct HC as [junk HC]. by inversion HC. + - destruct Hcont as (_ & _ & HC). + rewrite Hpvs in HC. by inversion HC. } + (* We commit. *) + pose (new_elts := map (get_value slots ({[i]} ∪ deqs)) new_pref ++ rest). + pose (new_pvs := proph_data sz ({[i]} ∪ deqs) rs'). + iMod "AU" as (elts_AU) "[He◯ [_ Hclose]]". + iDestruct (sync_elts with "He● He◯") as %<-. + iMod (update_elts _ _ _ new_elts with "He● He◯") as "[He● He◯]". + iMod ("Hclose" $! li new_elts with "[$He◯]") as "HΦ". + { iPureIntro. rewrite /new_elts /=. by rewrite /get_value Hslots_i. } + iModIntro. iSplitR "HΦ Hback_snap". + { pose (new_deqs := {[i]} ∪ deqs). + iNext. iExists back', new_pvs, new_pref, rest, cont', slots, new_deqs. + subst new_deqs. iFrame. iSplitL "Hℓ_ar". + { rewrite array_content_dequeue; [ done | by lia | done ]. } + iSplitL "Hp". + { iExists rs'. by iFrame "Hp". } + iPureIntro. repeat split_and; try done. + - intros k. split; intros Hk; first by apply Hstate. + intros Hk_in_deqs. apply elem_of_union in Hk_in_deqs. + destruct Hk_in_deqs as [Hk_is_i|Hk_in_deqs]. + + apply elem_of_singleton in Hk_is_i. subst k. + rewrite /array_get Hslots_i decide_False in Hi; last done. + rewrite /physical_value in Hi. rewrite Hslots_i in Hk. + inversion Hk; subst dw. inversion Hi. + + apply Hdeqs in Hk_in_deqs as (HContra & _). + rewrite HContra in Hk. inversion Hk. + - intros k Hk. + assert (k ∈ i :: new_pref) as HH%Hpref by set_solver +Hk. + destruct HH as (H1 & H2 & H3). repeat split; try done. + apply not_elem_of_union. split; last done. + apply not_elem_of_singleton. intros ->. + destruct cont' as [i1' i2'|bs]. + + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & [junk HC6]). + rewrite HC6 in Hpvs_ND. + apply NoDup_app in Hpvs_ND as (HND & _ & _). + apply NoDup_app in HND as (HND & _ & _). + apply NoDup_app in HND as (HND & _ & _). + apply NoDup_cons in HND as (HND & _). apply HND, Hk. + + destruct Hcont as (HC1 & HC2 & HC3). rewrite HC3 in Hpvs_ND. + apply NoDup_app in Hpvs_ND as (HND & _ & _). + apply NoDup_app in HND as (HND & _ & _). + apply NoDup_cons in HND as (HND & _). apply HND, Hk. + - intros k Hk. apply elem_of_union in Hk as [Hk%elem_of_singleton|Hk]. + + subst k. rewrite Hslots_i /=. + assert (dw = true) as ->. + { rewrite /array_get Hslots_i decide_False in Hi; last done. + rewrite /physical_value in Hi. destruct dw; first done. by inversion Hi. } + repeat split_and; [ done | by f_equal | .. ]. + rewrite /array_get Hslots_i decide_True; [ done | by set_solver ]. + + destruct (Hdeqs k Hk) as (H1 & H2 & H3). repeat split_and; try done. + rewrite /array_get. destruct (slots !! k) as [[[lk sk] wk]|]; last done. + rewrite decide_True; first done. by set_solver +Hk. + - by apply proph_data_NoDup. + - intros k Hk. by eapply (proph_data_sz sz _ _ _ Hk). + - destruct cont' as [i1' i2'|bs]. + + destruct Hcont as (((HC1 & HC2) & HC3) & HC4 & HC5 & HC6 & HC7 & HC8). + assert (i1' ≠ i) as Hi1'_not_i. + { intros ->. assert (i ∈ i :: new_pref) as Hpref_i%Hpref by set_solver. + by destruct Hpref_i as (_ & _ & Hpref_i). } + repeat split_and; try done. + * apply not_elem_of_union. split; last done. by apply not_elem_of_singleton. + * rewrite /array_get. destruct (slots !! i1') as [di1'|]; last by inversion HC2. + destruct di1' as [[li1' si1'] wi1']. rewrite decide_False; last by set_solver. + inversion HC5; subst wi1'. done. + * rewrite Hpvs /= in HC8. rewrite /new_pvs. by eapply prefix_cons_inv_2. + + destruct Hcont as (HC1 & HC2 & HC3). + repeat split_and; try done. rewrite Hpvs /= in HC3. by inversion HC3. } + wp_pures. done. } + (* If the enqueue at index [i] was not committed: contradiction. *) + exfalso. + assert (was_committed <$> slots !! i = Some false) as Hcom_i. + { rewrite Hslots_i. simpl. by f_equal. } + apply Hstate in Hcom_i. rewrite Hslots_i in Hcom_i. + inversion Hcom_i; subst dw. rewrite /array_get Hslots_i /= in Hi. + destruct (decide (i ∈ deqs)); by inversion Hi. + * (* The CmpXchg failed, we continue looping. *) + assert (array_content sz slots deqs !! i = Some NONEV). + { rewrite array_content_lookup; last by lia. by rewrite Hi. } + wp_apply (wp_cmpxchg_fail_offset _ _ _ _ _ (array_get slots deqs i) with "Hℓ_ar"); + [ by rewrite Hi | by rewrite Hi | by right | iIntros "Hℓa" (rs' ->) "Hp"]. + (* We can close the invariant. *) + iModIntro. iSplitR "AU Hback_snap Hi2_lower_bound". + { iNext. iExists _, _, _, _, cont', _, _. iFrame. iSplit; last done. + iExists rs'. rewrite Hpvs /= decide_True; last by lia. by iFrame. } + (* And conclude using the loop induction hypothesis. *) + wp_pures. assert (S n - 1 = n%nat) as -> by lia. iClear "Hval_wit_i". + iApply ("IH_loop" with "[] [] AU Hback_snap"). + - iPureIntro. lia. + - iPureIntro. destruct cont as [i1 i2|bs]; last done. + apply le_lt_or_eq in Hcont_i1. destruct Hcont_i1 as [Hi1|Hi1]; first lia. + exfalso. destruct Hinitial_cont as [-> Hinitial_cont]. + destruct Hcont as (HC1 & HC2 & HC3 & HC4 & HC5 & HC6 & HC7). + assert (is_Some (slots !! i1)) as Hslots_i1. { apply Hslots. lia. } + destruct Hslots_i1 as [[[li1 si1] wi1] Hslots_i1]. + rewrite /array_get Hslots_i1 decide_False in HC5; last done. + simpl in HC5. destruct wi1; last done. clear HC5. + rewrite array_content_lookup in H; last by lia. + rewrite /array_get in H. subst i i1. rewrite Hslots_i1 in H. + rewrite decide_False in H; last done. inversion H. + - by iFrame. +Qed. + +End herlihy_wing_queue. + +(** * Instantiation of the specification ***********************************) + +Definition atomic_cinc `{!heapG Σ, !savedPropG Σ, !hwqG Σ} : + spec.atomic_hwq Σ := + {| spec.new_queue_spec := new_queue_spec; + spec.enqueue_spec := enqueue_spec; + spec.dequeue_spec := dequeue_spec; + spec.hwq_content_exclusive := hwq_cont_exclusive |}. + +Typeclasses Opaque hwq_content is_hwq. diff --git a/theories/logatom/herlihy_wing_queue/spec.v b/theories/logatom/herlihy_wing_queue/spec.v new file mode 100644 index 0000000000000000000000000000000000000000..3145fe85844851a73fae074decf73ab4c6a62bcb --- /dev/null +++ b/theories/logatom/herlihy_wing_queue/spec.v @@ -0,0 +1,42 @@ +From iris.heap_lang Require Export lifting notation. +From iris.program_logic Require Export atomic. +Set Default Proof Using "Type". + +(** A general logically atomic interface for Herlihy-Wing queues. *) +Record atomic_hwq {Σ} `{!heapG Σ} := AtomicHWQ { + (* -- operations -- *) + new_queue : val; + enqueue : val; + dequeue : val; + (* -- other data -- *) + name : Type; + name_eqdec : EqDecision name; + name_countable : Countable name; + (* -- predicates -- *) + is_hwq (N : namespace) (sz : nat) (γ : name) (v : val) : iProp Σ; + hwq_content (γ : name) (ls : list loc) : iProp Σ; + (* -- predicate properties -- *) + is_hwq_persistent N sz γ v : Persistent (is_hwq N sz γ v); + hwq_content_timeless γ ls : Timeless (hwq_content γ ls); + hwq_content_exclusive γ ls1 ls2 : + hwq_content γ ls1 -∗ hwq_content γ ls2 -∗ False; + (* -- operation specs -- *) + new_queue_spec N (sz : nat) : + 0 < sz → + {{{ True }}} + new_queue #sz + {{{ v γ, RET v; is_hwq N sz γ v ∗ hwq_content γ [] }}}; + enqueue_spec N (sz : nat) (γ : name) (q : val) (l : loc) : + is_hwq N sz γ q -∗ + <<< ∀ (ls : list loc), hwq_content γ ls >>> + enqueue q #l @ ⊤ ∖ ↑N + <<< hwq_content γ (ls ++ [l]), RET #() >>>; + dequeue_spec N (sz : nat) (γ : name) (q : val) : + is_hwq N sz γ q -∗ + <<< ∀ (ls : list loc), hwq_content γ ls >>> + dequeue q @ ⊤ ∖ ↑N + <<< ∃ (l : loc) ls', ⌜ls = l :: ls'⌝ ∗ hwq_content γ ls', RET #l >>>; +}. +Arguments atomic_hwq _ {_}. + +Existing Instances is_hwq_persistent hwq_content_timeless. diff --git a/theories/logatom/lib/gc.v b/theories/logatom/lib/gc.v new file mode 100644 index 0000000000000000000000000000000000000000..9ed34a0be5ad0719889006d23b14fe4b47a3425f --- /dev/null +++ b/theories/logatom/lib/gc.v @@ -0,0 +1,243 @@ +From iris.algebra Require Import auth excl gmap. +From iris.base_logic.lib Require Export own invariants. +From iris.proofmode Require Import tactics. +From iris.heap_lang Require Export lang locations lifting. +Set Default Proof Using "Type". +Import uPred. + +Definition gcN: namespace := nroot .@ "gc". + +Definition gc_mapUR : ucmraT := gmapUR loc $ optionR $ exclR $ valO. + +Definition to_gc_map (gcm: gmap loc val) : gc_mapUR := (λ v : val, Excl' v) <$> gcm. + +Class gcG (Σ : gFunctors) := GcG { + gc_inG :> inG Σ (authR (gc_mapUR)); + gc_name : gname +}. + +Arguments gc_name {_} _ : assert. + +Class gcPreG (Σ : gFunctors) := { + gc_preG_inG :> inG Σ (authR (gc_mapUR)) +}. + +Definition gcΣ : gFunctors := + #[ GFunctor (authR (gc_mapUR)) ]. + +Instance subG_gcPreG {Σ} : subG gcΣ Σ → gcPreG Σ. +Proof. solve_inG. Qed. + +Section defs. + + Context `{!invG Σ, !heapG Σ, gG: gcG Σ}. + + Definition gc_inv_P: iProp Σ := + ((∃(gcm: gmap loc val), own (gc_name gG) (● to_gc_map gcm) ∗ ([∗ map] l ↦ v ∈ gcm, (l ↦ v)) ) )%I. + + Definition gc_inv : iProp Σ := inv gcN gc_inv_P. + + Definition gc_mapsto (l: loc) (v: val) : iProp Σ := own (gc_name gG) (◯ {[l := Excl' v]}). + + Definition is_gc_loc (l: loc) : iProp Σ := own (gc_name gG) (◯ {[l := None]}). + +End defs. + +Section to_gc_map. + + Lemma to_gc_map_valid gcm: ✓ to_gc_map gcm. + Proof. intros l. rewrite lookup_fmap. by case (gcm !! l). Qed. + + Lemma to_gc_map_empty: to_gc_map ∅ = ∅. + Proof. by rewrite /to_gc_map fmap_empty. Qed. + + Lemma to_gc_map_singleton l v: to_gc_map {[l := v]} = {[l := Excl' v]}. + Proof. by rewrite /to_gc_map fmap_insert fmap_empty. Qed. + + Lemma to_gc_map_insert l v gcm: + to_gc_map (<[l := v]> gcm) = <[l := Excl' v]> (to_gc_map gcm). + Proof. by rewrite /to_gc_map fmap_insert. Qed. + + Lemma to_gc_map_delete l gcm : + to_gc_map (delete l gcm) = delete l (to_gc_map gcm). + Proof. by rewrite /to_gc_map fmap_delete. Qed. + + Lemma lookup_to_gc_map_None gcm l : + gcm !! l = None → to_gc_map gcm !! l = None. + Proof. by rewrite /to_gc_map lookup_fmap=> ->. Qed. + + Lemma lookup_to_gc_map_Some gcm l v : + gcm !! l = Some v → to_gc_map gcm !! l = Some (Excl' v). + Proof. by rewrite /to_gc_map lookup_fmap=> ->. Qed. + + + Lemma lookup_to_gc_map_Some_2 gcm l w : + to_gc_map gcm !! l = Some w → ∃ v, gcm !! l = Some v. + Proof. + rewrite /to_gc_map lookup_fmap. rewrite fmap_Some. + intros (x & Hsome & Heq). eauto. + Qed. + + Lemma lookup_to_gc_map_Some_3 gcm l w : + to_gc_map gcm !! l = Some (Excl' w) → gcm !! l = Some w. + Proof. + rewrite /to_gc_map lookup_fmap. rewrite fmap_Some. + intros (x & Hsome & Heq). by inversion Heq. + Qed. + + Lemma excl_option_included (v: val) y: + ✓ y → Excl' v ≼ y → y = Excl' v. + Proof. + intros ? H. destruct y. + - apply Some_included_exclusive in H;[ | apply _ | done ]. + setoid_rewrite leibniz_equiv_iff in H. + by rewrite H. + - apply is_Some_included in H. + + by inversion H. + + by eapply mk_is_Some. + Qed. + + Lemma gc_map_singleton_included gcm l v : + {[l := Some (Excl v)]} ≼ to_gc_map gcm → gcm !! l = Some v. + Proof. + rewrite singleton_included. + setoid_rewrite Some_included_total. + intros (y & Hsome & Hincluded). + pose proof (lookup_valid_Some _ _ _ (to_gc_map_valid gcm) Hsome) as Hvalid. + pose proof (excl_option_included _ _ Hvalid Hincluded) as Heq. + rewrite Heq in Hsome. + apply lookup_to_gc_map_Some_3. + by setoid_rewrite leibniz_equiv_iff in Hsome. + Qed. +End to_gc_map. + +Lemma gc_init `{!invG Σ, !heapG Σ, !gcPreG Σ} E: + (|==> ∃ _ : gcG Σ, |={E}=> gc_inv)%I. +Proof. + iMod (own_alloc (● (to_gc_map ∅))) as (γ) "H●". + { rewrite auth_auth_valid. exact: to_gc_map_valid. } + iModIntro. + iExists (GcG Σ _ γ). + iAssert (gc_inv_P (gG := GcG Σ _ γ)) with "[H●]" as "P". + { + iExists _. iFrame. + by iApply big_sepM_empty. + } + iMod ((inv_alloc gcN E gc_inv_P) with "P") as "#InvGC". + iModIntro. iFrame "#". +Qed. + +Section gc. + Context `{!invG Σ, !heapG Σ, !gcG Σ}. + + Global Instance is_gc_loc_persistent (l: loc): Persistent (is_gc_loc l). + Proof. rewrite /is_gc_loc. apply _. Qed. + + Global Instance is_gc_loc_timeless (l: loc): Timeless (is_gc_loc l). + Proof. rewrite /is_gc_loc. apply _. Qed. + + Global Instance gc_mapsto_timeless (l: loc) (v: val): Timeless (gc_mapsto l v). + Proof. rewrite /is_gc_loc. apply _. Qed. + + Global Instance gc_inv_P_timeless: Timeless gc_inv_P. + Proof. rewrite /gc_inv_P. apply _. Qed. + + Lemma make_gc l v E: ↑gcN ⊆ E → gc_inv -∗ l ↦ v ={E}=∗ gc_mapsto l v. + Proof. + iIntros (HN) "#Hinv Hl". + iMod (inv_open_timeless _ gcN _ with "Hinv") as "[P Hclose]"=>//. + iDestruct "P" as (gcm) "[H● HsepM]". + destruct (gcm !! l) as [v' | ] eqn: Hlookup. + - (* auth map contains l --> contradiction *) + iDestruct (big_sepM_lookup with "HsepM") as "Hl'"=>//. + by iDestruct (mapsto_valid_2 with "Hl Hl'") as %?. + - iMod (own_update with "H●") as "[H● H◯]". + { + apply lookup_to_gc_map_None in Hlookup. + apply (auth_update_alloc _ (to_gc_map (<[l := v]> gcm)) (to_gc_map ({[l := v]}))). + rewrite to_gc_map_insert to_gc_map_singleton. + pose proof (to_gc_map_valid gcm). + setoid_rewrite alloc_singleton_local_update=>//. + } + iMod ("Hclose" with "[H● HsepM Hl]"). + + iExists _. + iDestruct (big_sepM_insert with "[HsepM Hl]") as "HsepM"=>//; iFrame. iFrame. + + iModIntro. by rewrite /gc_mapsto to_gc_map_singleton. + Qed. + + Lemma gc_is_gc l v: gc_mapsto l v -∗ is_gc_loc l. + Proof. + iIntros "Hgc_l". rewrite /gc_mapsto. + assert (Excl' v = (Excl' v) ⋅ None)%I as ->. { done. } + rewrite -op_singleton auth_frag_op own_op. + iDestruct "Hgc_l" as "[_ H◯_none]". + iFrame. + Qed. + + Lemma is_gc_lookup_Some l gcm: is_gc_loc l -∗ own (gc_name _) (● to_gc_map gcm) -∗ ⌜∃ v, gcm !! l = Some v⌝. + iIntros "Hgc_l H◯". + iCombine "H◯ Hgc_l" as "Hcomb". + iDestruct (own_valid with "Hcomb") as %Hvalid. + iPureIntro. + apply auth_both_valid in Hvalid as [Hincluded Hvalid]. + setoid_rewrite singleton_included in Hincluded. + destruct Hincluded as (y & Hsome & _). + eapply lookup_to_gc_map_Some_2. + by apply leibniz_equiv_iff in Hsome. + Qed. + + Lemma gc_mapsto_lookup_Some l v gcm: gc_mapsto l v -∗ own (gc_name _) (● to_gc_map gcm) -∗ ⌜ gcm !! l = Some v ⌝. + Proof. + iIntros "Hgc_l H●". + iCombine "H● Hgc_l" as "Hcomb". + iDestruct (own_valid with "Hcomb") as %Hvalid. + iPureIntro. + apply auth_both_valid in Hvalid as [Hincluded Hvalid]. + by apply gc_map_singleton_included. + Qed. + + (** An accessor to make use of [gc_mapsto]. + This opens the invariant *before* consuming [gc_mapsto] so that you can use + this before opening an atomic update that provides [gc_mapsto]!. *) + Lemma gc_access E: + ↑gcN ⊆ E → + gc_inv ={E, E ∖ ↑gcN}=∗ ∀ l v, gc_mapsto l v -∗ + (l ↦ v ∗ (∀ w, l ↦ w ==∗ gc_mapsto l w ∗ |={E ∖ ↑gcN, E}=> True)). + Proof. + iIntros (HN) "#Hinv". + iMod (inv_open_timeless _ gcN _ with "Hinv") as "[P Hclose]"=>//. + iIntros "!>" (l v) "Hgc_l". + iDestruct "P" as (gcm) "[H● HsepM]". + iDestruct (gc_mapsto_lookup_Some with "Hgc_l H●") as %Hsome. + iDestruct (big_sepM_delete with "HsepM") as "[Hl HsepM]"=>//. + iFrame. iIntros (w) "Hl". + iMod (own_update_2 with "H● Hgc_l") as "[H● H◯]". + { apply (auth_update _ _ (<[l := Excl' w]> (to_gc_map gcm)) {[l := Excl' w]}). + eapply singleton_local_update. + { by apply lookup_to_gc_map_Some in Hsome. } + by apply option_local_update, exclusive_local_update. + } + iDestruct (big_sepM_insert with "[Hl HsepM]") as "HsepM"; [ | iFrame | ]. + { apply lookup_delete. } + rewrite insert_delete. rewrite <- to_gc_map_insert. + iModIntro. iFrame. + iMod ("Hclose" with "[H● HsepM]"); [ iExists _; by iFrame | by iModIntro]. + Qed. + + Lemma is_gc_access l E: ↑gcN ⊆ E → gc_inv -∗ is_gc_loc l ={E, E ∖ ↑gcN}=∗ ∃ v, l ↦ v ∗ (l ↦ v ={E ∖ ↑gcN, E}=∗ ⌜True⌝). + Proof. + iIntros (HN) "#Hinv Hgc_l". + iMod (inv_open_timeless _ gcN _ with "Hinv") as "[P Hclose]"=>//. + iModIntro. + iDestruct "P" as (gcm) "[H● HsepM]". + iDestruct (is_gc_lookup_Some with "Hgc_l H●") as %Hsome. + destruct Hsome as [v Hsome]. + iDestruct (big_sepM_lookup_acc with "HsepM") as "[Hl HsepM]"=>//. + iExists _. iFrame. + iIntros "Hl". + iMod ("Hclose" with "[H● HsepM Hl]"); last done. + iExists _. iFrame. + by (iApply ("HsepM" with "Hl")). + Qed. + +End gc. diff --git a/theories/logatom/proph_erasure.v b/theories/logatom/proph_erasure.v new file mode 100644 index 0000000000000000000000000000000000000000..a46bc708c9f2e3a16ad832f18ee4dcaaffa75f97 --- /dev/null +++ b/theories/logatom/proph_erasure.v @@ -0,0 +1,862 @@ +From iris.heap_lang Require Export lang notation tactics. +From iris.program_logic Require Export adequacy. + +(** This file contains the proof that prophecies can be safely erased +from programs. We erase a program by replacing prophecy identifiers +with the unit values and respectively adapt the NewProph and Resolve +expressions. We prove that if a program e is safe then so is the +erasure of e. *) + + +Fixpoint erase_expr (e : expr) : expr := + match e with + | Val v => Val (erase_val v) + | Var x => Var x + | Rec f x e => Rec f x (erase_expr e) + | App e1 e2 => App (erase_expr e1) (erase_expr e2) + | UnOp op e => UnOp op (erase_expr e) + | BinOp op e1 e2 => BinOp op (erase_expr e1) (erase_expr e2) + | If e0 e1 e2 => If (erase_expr e0) (erase_expr e1) (erase_expr e2) + | Pair e1 e2 => Pair (erase_expr e1) (erase_expr e2) + | Fst e => Fst (erase_expr e) + | Snd e => Snd (erase_expr e) + | InjL e => InjL (erase_expr e) + | InjR e => InjR (erase_expr e) + | Case e0 e1 e2 => Case (erase_expr e0) (erase_expr e1) (erase_expr e2) + | Fork e => Fork (erase_expr e) + | AllocN e1 e2 => AllocN (erase_expr e1) (erase_expr e2) + | Load e => Load (erase_expr e) + | Store e1 e2 => Store (erase_expr e1) (erase_expr e2) + | CmpXchg e0 e1 e2 => CmpXchg (erase_expr e0) (erase_expr e1) (erase_expr e2) + | FAA e1 e2 => FAA (erase_expr e1) (erase_expr e2) + | NewProph => ((λ: <>, #LitErased)%V #()) + | Resolve e0 e1 e2 => Fst (Fst (erase_expr e0, erase_expr e1, erase_expr e2)) + end +with +erase_val (v : val) : val := + match v with + | LitV l => + LitV + match l with + | LitProphecy p => LitErased + | _ => l + end + | RecV f x e => RecV f x (erase_expr e) + | PairV v1 v2 => PairV (erase_val v1) (erase_val v2) + | InjLV v => InjLV (erase_val v) + | InjRV v => InjRV (erase_val v) + end. + +Lemma erase_expr_subst x v e : + erase_expr (subst x v e) = subst x (erase_val v) (erase_expr e) +with +erase_val_subst x v (w : val) : + erase_expr (subst x v w) = subst x (erase_val v) (erase_val w). +Proof. + - destruct e; simpl; try destruct decide; + rewrite ?erase_expr_subst ?erase_val_subst; auto. + - by destruct v; simpl. +Qed. + +Lemma erase_expr_subst' x v e : + erase_expr (subst' x v e) = subst' x (erase_val v) (erase_expr e). +Proof. destruct x; first done; apply erase_expr_subst. Qed. +Lemma erase_val_subst' x v (w : val) : + erase_expr (subst x v w) = subst x (erase_val v) (erase_val w). +Proof. destruct x; first done; apply erase_val_subst. Qed. + +Fixpoint erase_ectx_item (Ki : ectx_item) : list ectx_item := + match Ki with + | AppLCtx v2 => [AppLCtx (erase_val v2)] + | AppRCtx e1 => [AppRCtx (erase_expr e1)] + | UnOpCtx op => [UnOpCtx op] + | BinOpLCtx op v2 => [BinOpLCtx op (erase_val v2)] + | BinOpRCtx op e1 => [BinOpRCtx op (erase_expr e1)] + | IfCtx e1 e2 => [IfCtx (erase_expr e1) (erase_expr e2)] + | PairLCtx v2 => [PairLCtx (erase_val v2)] + | PairRCtx e1 => [PairRCtx (erase_expr e1)] + | FstCtx => [FstCtx] + | SndCtx => [SndCtx] + | InjLCtx => [InjLCtx] + | InjRCtx => [InjRCtx] + | CaseCtx e1 e2 => [CaseCtx (erase_expr e1) (erase_expr e2)] + | AllocNLCtx v2 => [AllocNLCtx (erase_val v2)] + | AllocNRCtx e1 => [AllocNRCtx (erase_expr e1)] + | LoadCtx => [LoadCtx] + | StoreLCtx v2 => [StoreLCtx (erase_val v2)] + | StoreRCtx e1 => [StoreRCtx (erase_expr e1)] + | CmpXchgLCtx v1 v2 => [CmpXchgLCtx (erase_val v1) (erase_val v2)] + | CmpXchgMCtx e0 v2 => [CmpXchgMCtx (erase_expr e0) (erase_val v2)] + | CmpXchgRCtx e0 e1 => [CmpXchgRCtx (erase_expr e0) (erase_expr e1)] + | FaaLCtx v2 => [FaaLCtx (erase_val v2)] + | FaaRCtx e1 => [FaaRCtx (erase_expr e1)] + | ResolveLCtx ctx v1 v2 => + erase_ectx_item ctx ++ + [PairLCtx (erase_val v1); PairLCtx (erase_val v2); FstCtx; FstCtx] + | ResolveMCtx e0 v2 => + [PairRCtx (erase_expr e0); PairLCtx (erase_val v2); FstCtx; FstCtx] + | ResolveRCtx e0 e1 => + [PairRCtx (erase_expr e0, erase_expr e1); FstCtx; FstCtx] + end. + +Definition erase_ectx (K : ectx heap_ectx_lang) : ectx heap_ectx_lang := + foldr (λ Ki K, erase_ectx_item Ki ++ K) [] K. + +Definition erase_tp (tp : list expr) : list expr := erase_expr <$> tp. + +Definition erase_heap (h : gmap loc val) : gmap loc val := erase_val <$> h. + +Definition erase_state (σ : state) : state := + {| heap := erase_heap (heap σ); used_proph_id := ∅ |}. + +Definition erase_cfg (ρ : cfg heap_lang) : cfg heap_lang := + (erase_tp ρ.1, erase_state ρ.2). + +Lemma erase_to_val e v : + to_val (erase_expr e) = Some v → ∃ v', to_val e = Some v' ∧ erase_val v' = v. +Proof. destruct e; intros ?; simplify_eq/=; eauto. Qed. + +Lemma erase_not_val e : to_val e = None → to_val (erase_expr e) = None. +Proof. by destruct e. Qed. + +Lemma erase_ectx_app K K' : erase_ectx (K ++ K') = erase_ectx K ++ erase_ectx K'. +Proof. + revert K'; induction K as [|Ki K IHK]; intros K'; simpl in *; first done. + by rewrite IHK assoc. +Qed. + +Lemma erase_ectx_expr K e : + erase_expr (fill K e) = fill (erase_ectx K) (erase_expr e). +Proof. + revert e; induction K as [|Ki K IHK] using rev_ind; intros e; + simpl in *; rewrite ?erase_ectx_app ?fill_app /=; first done. + induction Ki; + rewrite /= ?IHK ?IHKi ?fill_app //=. +Qed. + +(* helper lemma consider moving? *) +Lemma ectx_prim_step (K : list ectx_item) e1 σ1 κ e2 σ2 ef : + prim_step e1 σ1 κ e2 σ2 ef → prim_step (fill K e1) σ1 κ (fill K e2) σ2 ef. +Proof. + inversion 1; simpl in *; simplify_eq. + by rewrite -!fill_app; econstructor. +Qed. + +(* helper lemma consider moving? *) +Lemma rtc_congruence {A B} (f : A → B) (R : A → A → Prop) (R' : B → B → Prop) x y : + (∀ x y, R x y → R' (f x) (f y)) → rtc R x y → rtc R' (f x) (f y). +Proof. + intros Hcng. + induction 1 as [|e1 em e2 ?]; simpl in *; first done. + etrans; last done. + by apply rtc_once, Hcng. +Qed. + +(* helper lemma consider moving? *) +Lemma rtc_Ectx_pure_step (K : list ectx_item) e1 e2 : + rtc pure_step e1 e2 → rtc pure_step (fill K e1) (fill K e2). +Proof. + apply rtc_congruence. + by intros; eapply pure_step_ctx; first typeclasses eauto. +Qed. + +(* helper lemma consider moving? *) +Definition is_safe e σ := + is_Some (to_val e) ∨ reducible e σ. + +(* helper lemma consider moving? *) +Lemma is_safe_under_ectx K e σ : is_safe (fill K e) σ → is_safe e σ. +Proof. + rewrite /is_safe /reducible /=. + intros [?|(?&?&?&?&Hstp)]; simpl in *; + first by left; eapply fill_val. + destruct (to_val e) eqn:Hval; first by left; eauto. + inversion Hstp as [K' e1' ? He]; simpl in *; simplify_eq. + edestruct (step_by_val K K' e e1'); eauto; simplify_eq. + rewrite -fill_comp in He; apply fill_inj in He; simplify_eq. + right; eauto 10 using Ectx_step. +Qed. + +(* helper lemma consider moving? *) +Lemma head_step_safe e σ κ e' σ' efs : head_step e σ κ e' σ' efs → is_safe e σ. +Proof. rewrite /is_safe /reducible /=. eauto 10 using head_prim_step. Qed. + +(* helper lemma consider moving? *) +Lemma prim_step_safe e σ κ e' σ' efs : prim_step e σ κ e' σ' efs → is_safe e σ. +Proof. rewrite /is_safe /reducible /=. eauto 10. Qed. + +Lemma val_is_unboxed_erased v : + val_is_unboxed (erase_val v) ↔ val_is_unboxed v. +Proof. + destruct v; rewrite /= /lit_is_unboxed; repeat (done || simpl; case_match). +Qed. +Lemma vals_compare_safe_erase v1 v2 : + vals_compare_safe (erase_val v1) (erase_val v2) ↔ + vals_compare_safe v1 v2. +Proof. rewrite /vals_compare_safe !val_is_unboxed_erased. done. Qed. +Lemma vals_compare v1 v2 : + vals_compare_safe v1 v2 → + (v1 = v2) ↔ (erase_val v1 = erase_val v2). +Proof. + destruct v1, v2; rewrite /= /lit_is_unboxed; + repeat (done || (by intros [[] | []]) || simpl; case_match). +Qed. +(* Rewrite with [vals_compare] does not work, so derive a version +that wraps it in [bool_decide], that can be rewritten. *) +Lemma vals_compare_bdec v1 v2 : + vals_compare_safe v1 v2 → + bool_decide (v1 = v2) = bool_decide (erase_val v1 = erase_val v2). +Proof. + intros ?. eapply bool_decide_iff. apply vals_compare. done. +Qed. + +(** if un_op_eval succeeds on erased value, + the so should it on the original value. *) +Lemma un_op_eval_erase op v v' : + un_op_eval op (erase_val v) = Some v' ↔ + ∃ w, un_op_eval op v = Some w ∧ erase_val w = v'. +Proof. + destruct op; simpl; repeat case_match; + (split; [intros ?|intros [? [? ?]]]); + simplify_eq; simpl in *; simplify_eq; eauto. +Qed. + +(** if unbin_op_eval succeeds on erased value, + then so should it on the original value. *) +Lemma bin_op_eval_erase op v1 v2 v' : + bin_op_eval op (erase_val v1) (erase_val v2) = Some v' ↔ + ∃ w, bin_op_eval op v1 v2 = Some w ∧ erase_val w = v'. +Proof. + rewrite /bin_op_eval /bin_op_eval_int /bin_op_eval_bool; + destruct decide; simpl; repeat case_match; + (split; [intros ?|intros [? [? ?]]]); + simplify_eq; simpl in *; simplify_eq; eauto. + - eexists _; split; eauto; simpl. + erewrite bool_decide_iff; first by eauto. + apply vals_compare. done. + - repeat f_equal. + erewrite bool_decide_iff; first by eauto. + symmetry. apply vals_compare. done. + - exfalso. match goal with H: ¬_ |- _ => apply H end. + eapply vals_compare_safe_erase. done. + - exfalso. match goal with H: ¬_ |- _ => apply H end. + eapply vals_compare_safe_erase. done. +Qed. + +Lemma erase_heap_lookup h l : (erase_heap h) !! l = None ↔ h !! l = None. +Proof. rewrite lookup_fmap; by destruct (h !! l). Qed. + +Lemma erase_heap_lookup' h l : (erase_heap h) !! l = erase_val <$> h !! l. +Proof. by rewrite lookup_fmap. Qed. + +Lemma erase_heap_insert h l v : + erase_heap (<[l := v]> h) = <[l := erase_val v]> (erase_heap h). +Proof. + by rewrite /erase_heap fmap_insert. +Qed. + +Lemma fmap_heap_array f l vs : + f <$> heap_array l vs = heap_array l (f <$> vs). +Proof. + revert l; induction vs as [|v vs IHvs]; intros l; + first by rewrite /= fmap_empty. + by rewrite /= -!insert_union_singleton_l !fmap_insert IHvs. +Qed. + +Lemma erase_heap_array l n v σ : + erase_heap (heap_array l (replicate (Z.to_nat n) v) ∪ heap σ) = + heap_array l (replicate (Z.to_nat n) (erase_val v)) ∪ (erase_heap (heap σ)). +Proof. + apply map_eq => i. + rewrite /erase_heap lookup_fmap !lookup_union -fmap_replicate + -fmap_heap_array !lookup_fmap. + by destruct (heap_array l (replicate (Z.to_nat n) v) !! i); + destruct (heap σ !! i). +Qed. + +Lemma erase_state_init l n v σ: + erase_state (state_init_heap l n v σ) = + state_init_heap l n (erase_val v) (erase_state σ). +Proof. by rewrite /erase_state /state_init_heap /= erase_heap_array. Qed. + +Ltac simplify_erasure := + let is_constructor A := + let HF := fresh in + pose (match A with + | _ => A + end) as HF; + (simpl in HF); unify HF A; clear HF + in + lazymatch goal with + | H : erase_expr ?A = Val ?B |- _ => destruct A + | H : Val ?B = erase_expr ?A |- _ => destruct A + | H : _ = erase_val ?A |- _ => destruct A + | H : erase_val ?A = _ |- _ => destruct A + | H : erase_expr ?A = ?B |- _ => + lazymatch B with + | fill _ _ => fail + | context [fill _ _] => destruct A + | _ => is_constructor B; destruct A + end + | H : ?B = erase_expr ?A |- _ => + lazymatch B with + | fill _ _ => fail + | context [fill _ _] => destruct A + | _ => is_constructor B; destruct A + end + | H : Val _ = fill ?K ?e |- _ => + edestruct (to_val_fill_some K e); + first (by rewrite -H; eauto) + | H : head_step (Val _) _ _ _ _ _ |- _ => apply val_head_stuck in H + | H : Rec _ _ _ = fill ?K _ |- _ => + destruct K as [|[] ? _] using rev_ind; simpl in *; + rewrite ?fill_app in H + | H : un_op_eval _ (erase_val _) = Some _ |- _ => + apply un_op_eval_erase in H as [? [? ?]] + | H : bin_op_eval _ (erase_val _) (erase_val _) = Some _ |- _ => + apply bin_op_eval_erase in H as [? [? ?]] + | H : context [erase_heap (heap _) !! _] |- _ => + rewrite /erase_heap lookup_fmap in H + | H : _ <$> (?B !! ?C) = Some _ |- _ => + let Hf := fresh in destruct (B !! C) eqn:Hf; simpl in H; try done + | H : is_Some (_ <$> (?B !! ?C)) |- _ => destruct H as [? ?] + end; simplify_eq/=. + +Definition head_steps_to_erasure_of (e1 : expr) (σ1 : state) (e2 : expr) + (σ2 : state) (efs : list expr) := + ∃ κ' e2' σ2' efs', + head_step e1 σ1 κ' e2' σ2' efs' ∧ + erase_expr e2' = e2 ∧ erase_state σ2' = σ2 ∧ erase_tp efs' = efs. + +(** When the erased program makes a head step, so does the original program. *) +Lemma erased_head_step_head_step e1 σ1 κ e2 σ2 efs: + head_step (erase_expr e1) (erase_state σ1) κ e2 σ2 efs + → head_steps_to_erasure_of e1 σ1 e2 σ2 efs. +Proof. + intros Hstp; rewrite /head_steps_to_erasure_of. + inversion Hstp; + (repeat ((repeat simplify_erasure); try case_match; simplify_eq)); + try + (by (eexists _, _, _, _; simpl; split; + first match goal with + | |- head_step NewProph _ _ _ _ _ => by apply new_proph_id_fresh + | _ => by econstructor; + eauto using erase_heap_lookup + end; + try rewrite -val_for_compare_erase; + rewrite ?erase_expr_subst' /erase_state ?erase_heap_insert /=; + eauto using erase_state_init)); [| |]. + - (* case of allocN *) + eexists _, _, _, _; simpl; split; + first econstructor; try rewrite -val_for_compare_erase; + try setoid_rewrite <- erase_heap_lookup; + rewrite ?erase_heap_insert /=; eauto using erase_state_init. + - (* case of CmpXchg succeeding *) + match goal with + | H : bool_decide (erase_val _ = erase_val _) = _ |- _ => + rename H into Hvfc + end. + rewrite -vals_compare_bdec in Hvfc; last by eapply vals_compare_safe_erase. + eexists _, _, _, _; simpl; split. + { econstructor; eauto. rewrite -vals_compare_safe_erase //. } + rewrite Hvfc /erase_state ?erase_heap_insert /=; eauto. + - (* case of CmpXchg failing *) + match goal with + | H : bool_decide (erase_val _ = erase_val _) = _ |- _ => + rename H into Hvfc + end. + rewrite -vals_compare_bdec in Hvfc; last by eapply vals_compare_safe_erase. + eexists _, _, _, _; simpl; split. + { econstructor; eauto. rewrite -vals_compare_safe_erase //. } + rewrite Hvfc; eauto. +Qed. + +Lemma fill_to_resolve e0 (v1 v2 : val) K e' : + to_val e' = None → + Resolve e0 v1 v2 = fill K e' → + K = [] ∨ ∃ K' Ki, K = K' ++ [ResolveLCtx Ki v1 v2]. +Proof. + intros Hnv Hrs; simpl in *. + destruct K as [|Ki K _] using rev_ind; first by left. + rewrite fill_app in Hrs. + destruct Ki; simplify_eq/=; eauto; simplify_erasure. +Qed. + +(* helper lemma, move? *) +Lemma head_step_to_val e σ κ (v : val) σ' efs : + head_step e σ κ v σ' efs → + ∀ σ' κ' e' σ'' efs', head_step e σ' κ' e' σ'' efs' → is_Some (to_val e'). +Proof. + inversion 1; simplify_eq/= => ?????; inversion 1; simplify_eq; eauto; []. + match goal with | H : _ = ?A |- is_Some (to_val ?A) => rewrite -H end; eauto. +Qed. + +Ltac solve_for_trivial_pure_head_step := + let helper := + apply rtc_once; + apply pure_head_step_pure_step; + split; rewrite /head_reducible_no_obs /=; + first (by eauto using head_step); + intros ?????; inversion 1; simplify_eq; eauto + in + let go K e := + ((etrans; + first (apply (rtc_Ectx_pure_step K); helper)); simpl) + in + repeat match goal with + | |- rtc pure_step ?A _ => reshape_expr A go + | |- rtc pure_step ?A ?A => done + end. + +Lemma projs_pure_steps (v0 v1 v2 : val) : + rtc pure_step (Fst (Fst (v0, v1, v2))) v0. +Proof. solve_for_trivial_pure_head_step. Qed. + +Lemma prim_step_strip_ctx K K' e1 σ1 κ e2 σ2 efs κ' e2' σ2' efs' : + head_step e1 σ1 κ' e2' σ2' efs' → + prim_step (fill K (fill K' e1)) σ1 κ e2 σ2 efs → + ∃ e2'', e2 = fill K e2'' ∧ prim_step (fill K' e1) σ1 κ e2'' σ2 efs. +Proof. + rewrite -fill_app => Hhstp Hpstp. + inversion Hpstp as [K'' ?? HK'' ? hhstp']; simplify_eq/=. + edestruct (step_by_val (K' ++ K)) as [K3 HK3]; + eauto using val_head_stuck; simplify_eq/=. + rewrite !fill_app in HK''; repeat apply fill_inj in HK''; simplify_eq. + rewrite !fill_app. + eexists; split; first done. + rewrite -!fill_app. + by eapply Ectx_step'. +Qed. + +Lemma Resolve_3_vals_unsafe (v0 v1 v2 : val) σ : + is_safe (Resolve v0 v1 v2) σ → False. +Proof. + intros [[? ?]|(?&?&?&?&Hstp)]; simpl in *; first done. + inversion Hstp as [??? Hfill ? Hhstp']; simpl in *; simplify_eq. + destruct K as [|Ki K _] using rev_ind; simpl in *; simplify_eq. + - inversion Hhstp'; simplify_eq. + match goal with + | H : head_step (Val _) _ _ _ _ _ |- _ => + by apply val_head_stuck in H + end. + - clear Hstp; rewrite fill_app /= in Hfill. + apply val_head_stuck in Hhstp'. + destruct Ki; simpl in *; simplify_eq; + match goal with + | H : Val _ = fill_item ?Ki ?e |- _ => + by destruct (fill_item_val Ki e) as [? [-> ->]%to_val_fill_some]; + first by rewrite -H; eauto + | H : Val _ = fill ?K ?e |- _ => + edestruct (to_val_fill_some K e); + first (by rewrite -H; eauto); simplify_eq + end. +Qed. + +(** (e1, σ1) takes a [prim_step] to (e2', σ2') forking threads efs' + such that σ2 is the erasure of σ2' and efs is the erasure of + efs'. Furthermore, e2 takes [pure_steps] to match up with e2. It is + crucial for us that e2 takes [pure_step]s because we need to know + that e2 does not get stuck and that the steps are deterministic. + + Essentially, the main part of the erasure proof's argument is that if + the erased program takes steps, then the original program also takes + matching steps. This however, does not entirely hold. In cases where + the erasure of [Resovle] takes a step, the original program + immediately produces the value while the erased program has to still + perform projections [Fst] to get the result (see the [Resolve] case + of [erase_expr]). For this purpose, we prove that in those cases (and + also in general) the erased program also takes a number of (possibly + zero) steps so that the original and the erased programs are matched + up again. *) +Definition prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs := + ∃ e2' σ2' κ' efs' e2'', + prim_step e1 σ1 κ' e2' σ2' efs' ∧ rtc pure_step e2 e2'' ∧ + erase_expr e2' = e2'' ∧ erase_state σ2' = σ2 ∧ erase_tp efs' = efs. + +Lemma prim_step_matched_by_erased_steps_ectx K e1 σ1 e2 σ2 efs : + prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs → + prim_step_matched_by_erased_steps (fill K e1) σ1 (fill (erase_ectx K) e2) σ2 efs. +Proof. + intros (?&?&?&?&?&?&?&?&?&?); simplify_eq/=. + eexists _, _, _, _, _; repeat split. + - by eapply ectx_prim_step. + - by rewrite erase_ectx_expr; apply rtc_Ectx_pure_step. +Qed. + +Definition is_Resolve (e : expr) := + match e with | Resolve _ _ _ => True | _ => False end. + +Global Instance is_Resolve_dec e : Decision (is_Resolve e). +Proof. destruct e; solve_decision. Qed. + +Ltac solve_is_safe := + match goal with + | H : is_safe ?C ?B |- is_safe ?A ?B => + let tac K e := by (apply (is_safe_under_ectx K)) in + reshape_expr C tac + | H : head_step ?A ?B _ _ _ _ |- is_safe ?A ?B => by eapply head_step_safe + | H : prim_step ?A ?B _ _ _ _ |- is_safe ?A ?B => by eapply prim_step_safe + end. + +Ltac solve_prim_step_ectx := + match goal with + | |- prim_step ?A _ _ _ _ _ => + let tac K e := + lazymatch K with [] => fail | _ => (apply (ectx_prim_step K)) end + in + reshape_expr A tac; eauto + end. + +Ltac solve_rtc_pure_step_ectx := + match goal with + | |- rtc pure_step ?A _ => + let tac K e := + lazymatch K with [] => fail | _ => (apply (rtc_Ectx_pure_step K)) end + in + reshape_expr A tac; eauto + end. + +Lemma non_resolve_prim_step_matched_by_erased_steps_ectx_item + Ki e1 e1' σ1 e2 σ2 efs : + to_val e1' = None → + ¬ is_Resolve e1 → + is_safe e1 σ1 → + erase_expr e1 = fill_item Ki e1' → + (∀ e1, erase_expr e1 = e1' → is_safe e1 σ1 → + prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs) → + prim_step_matched_by_erased_steps e1 σ1 (fill_item Ki e2) σ2 efs. +Proof. + intros Hnv Hnr Hsf He1 IH. + destruct Ki; repeat simplify_erasure; try done; + match goal with + | |- prim_step_matched_by_erased_steps ?A _ _ _ _ => + let tac K e := + lazymatch K with + | [] => fail + | _ => apply (prim_step_matched_by_erased_steps_ectx K) + end + in + reshape_expr A tac; apply IH; [done|solve_is_safe] + end. +Qed. + +Lemma prim_step_matched_by_erased_steps_ectx_item Ki K e1 e1' σ1 e2 σ2 efs κ : + head_step e1' (erase_state σ1) κ e2 σ2 efs → + is_safe e1 σ1 → + erase_expr e1 = fill_item Ki (fill K e1') → + (∀ K' e1, length K' ≤ length K → + erase_expr e1 = (fill K' e1') → is_safe e1 σ1 → + prim_step_matched_by_erased_steps e1 σ1 (fill K' e2) σ2 efs) → + prim_step_matched_by_erased_steps e1 σ1 (fill_item Ki (fill K e2)) σ2 efs. +Proof. + intros Hhstp Hsf He1 IH; simpl in *. + (** Case split on whether e1 is a [Resolve] expression. *) + destruct (decide (is_Resolve e1)); last first. + { (** e1 is not a [Resolve] expression. *) + eapply non_resolve_prim_step_matched_by_erased_steps_ectx_item; eauto; [|]. + - by eapply fill_not_val, val_head_stuck. + - intros; eapply (IH K); simpl; eauto with lia. } + (** e1 is a [Resolve] expression. *) + destruct Ki; repeat simplify_erasure; try done; []. + destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|]. + { (* case where (Fst (erase_expr e1_1, erase_expr e1_2, erase_expr e1_3)) *) + (* takes a head_step; it is impossible! *) + inversion Hhstp. } + rewrite fill_app /= in He1. + destruct Ki; simplify_eq/=; rewrite fill_app /=. + destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|]. + { (* case where (erase_expr e1_1, erase_expr e1_2, erase_expr e1_3) *) + (* takes a head_step; it is impossible! *) + inversion Hhstp. } + rewrite fill_app /= in He1. + destruct Ki; simplify_eq/=; rewrite fill_app /=. + - destruct K as [|Ki K _] using rev_ind; simplify_eq/=; [|]. + { (** [Resolve v0 v1 v2] is not safe! *) + inversion Hhstp; repeat simplify_erasure. + exfalso; eauto using Resolve_3_vals_unsafe. } + rewrite fill_app /= in He1. + destruct Ki; simplify_eq/=; rewrite fill_app /=. + + (** e1 is of the form ([Resolve] e10 e11 v0) and e11 takes a prim_step. *) + destruct Hsf as [[? ?]| (?&?&?&?&Hrpstp)]; first done; simpl in *. + inversion Hrpstp as [??? Hrs ? Hhstp']; simpl in *; simplify_eq. + repeat simplify_erasure. + edestruct fill_to_resolve as [?|[K' [Ki HK]]]; eauto; + [by eapply val_head_stuck| |]; simplify_eq/=. + * (** e1 is of the form ([Resolve] e10 e11 v0) and e11 takes a head_step. *) + inversion Hhstp'; simplify_eq. + edestruct (IH K) as (?&?&?&?&?&Hpstp&?&?&?&?); + [rewrite !app_length /=; lia|done|by eapply head_step_safe|]; + simplify_eq/=. + apply head_reducible_prim_step in Hpstp; simpl in *; + last rewrite /head_reducible /=; eauto 10. + edestruct head_step_to_val as [? ?%of_to_val]; eauto; simplify_eq. + eexists _, _, _, _, _; repeat split; + first (by apply head_prim_step; econstructor; eauto); auto. + etrans. + { by apply (rtc_Ectx_pure_step + [PairLCtx _; PairLCtx _; FstCtx; FstCtx]). } + apply projs_pure_steps. + * (** e1 is of the form ([Resolve] e10 v v0) and e10 takes a + (non-head) prim_step. *) + rewrite fill_app in Hrs; simplify_eq/=. + edestruct (IH K) as (?&?&?&?&?&Hpstp&?&?&?&?); + [rewrite !app_length; lia|done| |]. + { change (fill_item Ki) with (fill [Ki]). + by rewrite -fill_app; eapply prim_step_safe, Ectx_step. } + simplify_eq/=. + eapply (prim_step_strip_ctx [_]) in Hpstp as [e2'' [He2'' Hpstp]]; + last done; simplify_eq/=. + eexists _, _, _, _, _; repeat split; eauto; + first by apply (ectx_prim_step [ResolveLCtx _ _ _]). + by apply (rtc_Ectx_pure_step [PairLCtx _; PairLCtx _; FstCtx; FstCtx]). + + (** e1 is of the form ([Resolve] e1_ e1_2 v) and e1_2 takes a prim_step. *) + repeat simplify_erasure. + apply (prim_step_matched_by_erased_steps_ectx [ResolveMCtx _ _]). + apply IH; [rewrite !app_length /=; lia|done|solve_is_safe]. + - (** e1 is of the form ([Resolve] e1_ e1_2 e13) and e1_3 takes a prim_step. *) + apply (prim_step_matched_by_erased_steps_ectx [ResolveRCtx _ _]). + apply IH; [rewrite !app_length /=; lia|done|solve_is_safe]. +Qed. + +Lemma erased_prim_step_prim_step e1 σ1 κ e2 σ2 efs: + prim_step (erase_expr e1) (erase_state σ1) κ e2 σ2 efs → is_safe e1 σ1 → + prim_step_matched_by_erased_steps e1 σ1 e2 σ2 efs. +Proof. + intros Hstp He1sf. + inversion Hstp as [K e1' e2' He1 ? Hhstp]; clear Hstp; + simpl in *; simplify_eq. + set (len := length K); assert (length K = len) as Hlen by done; clearbody len. + revert K Hlen e1 He1 He1sf. + induction len as [m IHm]using lt_wf_ind; intros K Hlen e1 He1 He1sf; + simplify_eq. + destruct K as [|Ki K _] using rev_ind; simpl in *; simplify_eq. + { apply erased_head_step_head_step in Hhstp as (?&?&?&?&?&<-&?&<-). + eexists _, _, _, _, _; repeat split; + first (by apply head_prim_step); auto using rtc_refl. } + rewrite app_length in IHm; simpl in *. + rewrite fill_app /=; rewrite fill_app /= in He1. + eapply prim_step_matched_by_erased_steps_ectx_item; eauto; []. + { intros; simpl in *; apply (IHm (length K')); auto with lia. } +Qed. + +Lemma head_step_erased_head_step e1 σ1 κ e2 σ2 ef: + head_step e1 σ1 κ e2 σ2 ef → + ∃ e2' σ2' ef', prim_step (erase_expr e1) (erase_state σ1) [] e2' σ2' ef'. +Proof. + induction 1 as [| | | | | | | | | | | | | | | | | | | |????????? IH]; + simpl in *; simplify_eq; + try (by do 3 eexists; apply head_prim_step; econstructor); + [| | | | | | |]. + - (* UnOp *) + do 3 eexists; apply head_prim_step; econstructor. + by apply un_op_eval_erase; eauto. + - (* BinOp *) + do 3 eexists; apply head_prim_step; econstructor. + by apply bin_op_eval_erase; eauto. + - (* AllocN *) + do 3 eexists; apply head_prim_step; econstructor; eauto. + intros; rewrite erase_heap_lookup; eauto. + - (* Load *) + do 3 eexists; apply head_prim_step; econstructor. + rewrite /erase_state /state_upd_heap /= erase_heap_lookup'. + by destruct lookup; simplify_eq. + - (* Store *) + match goal with | H : is_Some _ |- _ => inversion H end. + do 3 eexists; apply head_prim_step; econstructor. + by rewrite /erase_state /state_upd_heap /= erase_heap_lookup' H0; eauto. + - (* CmpXchg *) + match goal with + | H : vals_compare_safe ?A ?B |- _ => + destruct (bool_decide (A = B)) eqn:Heqvls + end. + + do 3 eexists; apply head_prim_step; + econstructor; last (by eauto); + last (by apply vals_compare_safe_erase); []. + match goal with + | H : heap _ !! _ = _ |- _ => + by rewrite /erase_state /state_upd_heap /= erase_heap_lookup' H + end. + + do 3 eexists; apply head_prim_step; + econstructor; last (by eauto); + last (by apply vals_compare_safe_erase); []. + match goal with + | H : heap _ !! _ = _ |- _ => + by rewrite /erase_state /state_upd_heap /= erase_heap_lookup' H + end. + - (* FAA *) + do 3 eexists; apply head_prim_step; econstructor. + by rewrite /erase_state /state_upd_heap /= erase_heap_lookup' H. + - (* Resolve *) + destruct IH as (?&?&?&?). + eexists _, _, _; + by apply (ectx_prim_step [PairLCtx _; PairLCtx _;FstCtx; FstCtx]). +Qed. + +Lemma reducible_erased_reducible e σ : + reducible e σ → reducible (erase_expr e) (erase_state σ). +Proof. + intros (?&?&?&?&Hpstp); simpl in *. + inversion Hpstp; simplify_eq/=. + rewrite erase_ectx_expr. + edestruct head_step_erased_head_step as (?&?&?&?); first done; simpl in *. + eexists _, _, _, _; eapply ectx_prim_step; eauto. +Qed. + +Definition erase_prop (φ : val → state → Prop) (v : val) (σ : state) : Prop := + ∃ v' σ', erase_val v' = v ∧ erase_state σ' = σ ∧ φ v' σ'. + +(* move to std++ perhaps? *) +Lemma nsteps_inv_r {A} (R : A → A → Prop) n x z : + nsteps R (S n) x z → ∃ y, nsteps R n x y ∧ R y z. +Proof. + revert x z; induction n as [|n IHn]; intros x z. + - inversion 1 as [|? ? ? ? ? Hstps]; inversion Hstps; simplify_eq. + by eexists; split; first constructor. + - inversion 1 as [|? ? ? ? ? Hstps]; simplify_eq. + apply IHn in Hstps as [w [Hstps Hw]]. + eexists; split; last done. + econstructor; eauto. +Qed. + +Definition pure_step_tp (t1 t2 : list expr) := Forall2 (rtc pure_step )t1 t2. + +Lemma rtc_pure_step_val v e : rtc pure_step (of_val v) e → to_val e = Some v. +Proof. + intros ?; rewrite <- to_of_val. + f_equal; symmetry; eapply rtc_nf; eauto. + intros [? [Hstp _]]. + by destruct (Hstp inhabitant) as (?&?&?&?%val_stuck). +Qed. + +Lemma pure_step_tp_safe t1 t2 e1 σ : + (∀ e2, e2 ∈ t2 → is_safe e2 σ) → pure_step_tp t1 (erase_tp t2) → + e1 ∈ t1 → is_safe e1 (erase_state σ). +Proof. + intros Ht2 Hpr [i He1]%elem_of_list_lookup_1. + eapply Forall2_lookup_l in Hpr as [e2' [He2' Hpr]]; simpl in *; eauto. + rewrite /erase_tp list_lookup_fmap in He2'. + destruct (t2 !! i) eqn:He2; simplify_eq/=. + apply elem_of_list_lookup_2, Ht2 in He2. + clear -Hpr He2. + inversion Hpr as [|??? [? _]]; simplify_eq. + - destruct He2 as [[? ?%of_to_val]|]; simplify_eq/=; first by left; eauto. + by right; apply reducible_erased_reducible. + - right; eauto using reducible_no_obs_reducible. +Qed. + +Lemma erased_step_pure_step_tp t1 σ1 t2 σ2 t3 : + erased_step (t1, σ1) (t2, σ2) → + pure_step_tp t1 t3 → + (σ1 = σ2 ∧ pure_step_tp t2 t3) ∨ + (∃ (i : nat) e t2' efs e' κ, t3 !! i = Some e ∧ t1 !! i = Some e ∧ + length t1 = length t2' ∧ t2 = t2' ++ efs ∧ + (∀ j, j ≠ i → t1 !! j = t2' !! j) ∧ + t2' !! i = Some e' ∧ prim_step e σ1 κ e' σ2 efs). +Proof. + intros [κ Hes] Hps; simpl in *. + inversion Hes as [e σ e' σ' ? t11 t12 ?? Hpstp]; simplify_eq/=. + edestruct @Forall2_lookup_l as [e'' [Hlu He'']]; + first exact Hps; first apply list_lookup_middle; first done. + inversion He'' as [|??? [_ Hprs]]; simplify_eq/=. + - right; eexists (length t11), e'', (t11 ++ e' :: t12), _, _, _; + repeat split; eauto; rewrite ?app_length -?app_assoc; + eauto using list_lookup_middle. + intros j Hneq. + destruct (decide (j < length t11)%nat). + + rewrite !lookup_app_l //. + + rewrite !lookup_app_r; try lia. + by destruct (j - length t11)%nat eqn:Heq; first lia. + - edestruct Hprs as (?&?&?&?); eauto; simplify_eq. + left; split; first done; rewrite app_nil_r. + apply Forall2_app_inv_l in Hps as (t31&t32&Hps1&Hps2&->). + apply Forall2_app; first done. + apply Forall2_cons_inv_l in Hps2 as (e'&t32'&?&?&?); simplify_eq. + apply Forall2_cons; last done. + rewrite list_lookup_middle in Hlu; last by eapply Forall2_length. + by simplify_eq. +Qed. + +(* move to std++ or find an easy way to prove this! *) +Lemma elem_of_list_split_length {A : Type} (l : list A) i (x : A) : + l !! i = Some x → ∃ l1 l2 : list A, l = l1 ++ x :: l2 ∧ i = length l1. +Proof. + revert i; induction l as [|a l IHl]; intros i Hl; first done. + destruct i; simplify_eq/=. + - exists []; eauto. + - destruct (IHl _ Hl) as (?&?&?&?); simplify_eq/=. + eexists (a :: _); eauto. +Qed. + +Lemma erased_steps_extend i t1 σ1 t2 σ2 (e : expr) κ e' σ3 efs : + rtc erased_step (t1, σ1) (t2, σ2) → + t2 !! i = Some e → + prim_step e σ2 κ e' σ3 efs → + rtc erased_step (t1, σ1) (<[i := e']>t2 ++ efs, σ3). +Proof. + intros Hes Hlu Hpstp; simpl in *. + etrans; first done. + apply rtc_once. + edestruct (elem_of_list_split_length t2) as (t21&t22&?&?); + first (by eauto using elem_of_list_lookup_2); simplify_eq. + do 2 econstructor; eauto. + replace (length t21) with (length t21 + 0)%nat by lia. + by rewrite insert_app_r /= -app_assoc /=. +Qed. + +Lemma erasure e σ φ : + adequate NotStuck e σ φ → + adequate NotStuck (erase_expr e) (erase_state σ) (erase_prop φ). +Proof. + intros Hade; simpl in *. + cut (∀ t2 σ2, + rtc erased_step ([erase_expr e], erase_state σ) (t2, σ2) → + (∃ t2' t2'' σ2', + rtc erased_step ([e], σ) (t2'', σ2') ∧ + t2' = erase_tp t2'' ∧ σ2 = erase_state σ2' ∧ + pure_step_tp t2 t2')). + { intros Hreach; split; simpl in *. + - intros ? ? ? Hrtc; edestruct (Hreach _ _ Hrtc) as + (t2'&t2''&σ2'&Hos&Ht2'&Hσ2&Hptp); simplify_eq/=. + apply Forall2_cons_inv_l in Hptp as (oe&t3&Hoe%rtc_pure_step_val&_&?); + destruct t2''; simplify_eq/=. + apply erase_to_val in Hoe as (?&?%of_to_val&?); simplify_eq. + pose proof (adequate_result _ _ _ _ Hade _ _ _ Hos); + rewrite /erase_prop; eauto. + - intros ? ? ? Hs Hrtc He2; edestruct (Hreach _ _ Hrtc) as + (t2'&t2''&σ2'&Hos&Ht2'&Hσ2&Hptp); simplify_eq/=. + eapply pure_step_tp_safe; [|done..]. + intros e2' He2'. + apply (adequate_not_stuck _ _ _ _ Hade _ _ _ eq_refl Hos He2'). } + intros t2 σ2 [n Hstps]%rtc_nsteps; simpl in *; revert t2 σ2 Hstps. + induction n. + { intros t2 σ2 Hstps; inversion Hstps; simplify_eq /=. + repeat econstructor. } + intros t2 σ2 Hstps. + apply nsteps_inv_r in Hstps as [[t3 σ3] [Hstps Hρ]]; simpl in *. + destruct (IHn _ _ Hstps) as (t2'&t2''&σ2'&Hostps&?&?&Hprstps); simplify_eq. + edestruct erased_step_pure_step_tp as [[? Hint]|Hext]; simplify_eq/=; + eauto 10; [|done..]. + destruct Hext as (i&ei&t2'&efs&e'&κ&Hi1&Hi2&Hlen&?&Hlu&Hi3&Hpstp); + simplify_eq. + rewrite /erase_tp list_lookup_fmap in Hi1. + destruct (t2'' !! i) as [eio|] eqn:Heq; simplify_eq/=. + edestruct erased_prim_step_prim_step as + (eio' & σ3 & κ' & efs' & ee & Heiopstp & Hprstps' & ?&?&?); first done; + last simplify_eq/=. + { eapply adequate_not_stuck; eauto using elem_of_list_lookup_2. } + eexists _, _, _; repeat split; first eapply erased_steps_extend; eauto. + rewrite /pure_step_tp /erase_tp fmap_app. + apply Forall2_app; last done. + apply Forall2_same_length_lookup; split. + { apply Forall2_length in Hprstps; rewrite fmap_length in Hprstps. + by rewrite fmap_length insert_length -Hlen. } + intros j x y. + destruct (decide (i = j)); simplify_eq. + { rewrite Hi3 list_lookup_fmap list_lookup_insert /=; + last by eapply lookup_lt_Some. + by intros ? ?; simplify_eq. } + rewrite list_lookup_fmap list_lookup_insert_ne // -list_lookup_fmap. + intros ? ?. + eapply Forall2_lookup_lr; eauto. + by rewrite Hlu. +Qed. diff --git a/theories/logatom/rdcss/rdcss.v b/theories/logatom/rdcss/rdcss.v new file mode 100644 index 0000000000000000000000000000000000000000..4b2f3e886b9f361fd246558ef0fe173807347ca8 --- /dev/null +++ b/theories/logatom/rdcss/rdcss.v @@ -0,0 +1,646 @@ +From iris.algebra Require Import excl auth agree frac list cmra csum. +From iris.base_logic.lib Require Export invariants. +From iris.program_logic Require Export atomic. +From iris.proofmode Require Import tactics. +From iris.heap_lang Require Import proofmode notation. +From iris_examples.logatom.rdcss Require Import spec. +Import uPred bi List Decidable. +Set Default Proof Using "Type". + +(** Using prophecy variables with helping: implementing a simplified version of + the restricted double-compare single-swap from "A Practical Multi-Word Compare-and-Swap Operation" by Harris et al. (DISC 2002) + *) + +(** * Implementation of the functions. *) + +(* 1) The M location l_m can differ when helping another thread in the same RDCSS instance + (Harris et al. refer to it as a location in the control section.) + 2) The N location l_n identifies a single RDCSS instance. + (Harris et al. refer to it as a location in the data section.) + 3) There are two kinds of values stored at N locations + + 3.1) The value is injL n for some n: no operation is on-going and the logical state is n. + 3.2) The value is injR l_descr for some location l_descr (which we call a descriptor): + l_descr must point to a tuple (l_m, m1, n1, n2, p) for some M location l_m, values m1, n1, n2 and prophecy p. + In this case an operation is on-going with corresponding M location l_m. +*) + +(* + new_rdcss(n) := ref (injL n) + *) +Definition new_rdcss : val := λ: "n", ref (InjL "n"). + +(* + complete(l_descr, l_n) := + let (l_m, m1, n1, n2, p) := !l_descr in + (* data = (l_m, m1, n1, n2, p) *) + let tid_ghost = NewProph in + let n_new = (if !l_m = m1 then n1 else n2) in + Resolve (CmpXchg l_n (InjR l_descr) (ref (InjL n_new))) p tid_ghost ; #(). + *) +Definition complete : val := + λ: "l_descr" "l_n", + let: "data" := !"l_descr" in + (* data = (l_m, m1, n1, n2, p) *) + let: "l_m" := Fst (Fst (Fst (Fst ("data")))) in + let: "m1" := Snd (Fst (Fst (Fst ("data")))) in + let: "n1" := Snd (Fst (Fst ("data"))) in + let: "n2" := Snd (Fst ("data")) in + let: "p" := Snd ("data") in + (* Create a thread identifier using NewProph. + tid_ghost is not used as a traditional prophecy, i.e. it is never resolved. + The reason we use NewProph is so that we can use the erasure theorem to argue that + it has no effect on the code. + *) + let: "tid_ghost" := NewProph in + let: "n_new" := (if: !"l_m" = "m1" then "n2" else "n1") in + Resolve (CmpXchg "l_n" (InjR "l_descr") (InjL "n_new")) "p" "tid_ghost" ;; #(). + +(* + get(l_n) := + match: !l_n with + | injL n => n + | injR (l_descr) => + complete(l_descr, l_n); + get(l_n) + end. + *) +Definition get : val := + rec: "get" "l_n" := + match: !"l_n" with + InjL "n" => "n" + | InjR "l_descr" => + complete "l_descr" "l_n" ;; + "get" "l_n" + end. + +(* + rdcss(l_m, l_n, m1, n1, n2) := + let p := NewProph in + let l_descr := ref (l_m, m1, n1, n2, p) in + (rec: rdcss_inner() + let (r, b) := CmpXchg(l_n, InjL n1, InjR l_descr) in + match r with + InjL n => + if b then + complete(l_descr, l_n) ; n1 + else + n + | InjR l_descr_other => + complete(l_descr_other, l_n) ; + rdcss_inner() + end + )() +*) +Definition rdcss: val := + λ: "l_m" "l_n" "m1" "n1" "n2", + (* allocate fresh descriptor *) + let: "p" := NewProph in + let: "l_descr" := ref ("l_m", "m1", "n1", "n2", "p") in + (* start rdcss computation with allocated descriptor *) + ( rec: "rdcss_inner" "_" := + let: "r" := CmpXchg "l_n" (InjL "n1") (InjR "l_descr") in + match: Fst "r" with + InjL "n" => + (* non-descriptor value read, check if CmpXchg was successful *) + if: Snd "r" then + (* CmpXchg was successful, finish operation *) + complete "l_descr" "l_n" ;; "n1" + else + (* CmpXchg failed, hence we could linearize at the CmpXchg *) + "n" + | InjR "l_descr_other" => + (* a descriptor from a different operation was read, try to help and then restart *) + complete "l_descr_other" "l_n" ;; + "rdcss_inner" #() + end + ) #(). + +(** ** Proof setup *) + +Definition valUR := authR $ optionUR $ exclR valO. +Definition tokenUR := exclR unitO. +Definition one_shotUR := csumR (exclR unitO) (agreeR unitO). + +Class rdcssG Σ := RDCSSG { + rdcss_valG :> inG Σ valUR; + rdcss_tokenG :> inG Σ tokenUR; + rdcss_one_shotG :> inG Σ one_shotUR; + }. + +Definition rdcssΣ : gFunctors := + #[GFunctor valUR; GFunctor tokenUR; GFunctor one_shotUR]. + +Instance subG_rdcssΣ {Σ} : subG rdcssΣ Σ → rdcssG Σ. +Proof. solve_inG. Qed. + +Section rdcss. + Context {Σ} `{!heapG Σ, !rdcssG Σ, !gcG Σ }. + Context (N : namespace). + + Local Definition descrN := N .@ "descr". + Local Definition rdcssN := N .@ "rdcss". + + (** Updating and synchronizing the value RAs *) + + Lemma sync_values γ_n (n m : val) : + own γ_n (● Excl' n) -∗ own γ_n (◯ Excl' m) -∗ ⌜n = m⌝. + Proof. + iIntros "H● H◯". iCombine "H●" "H◯" as "H". iDestruct (own_valid with "H") as "H". + by iDestruct "H" as %[H%Excl_included%leibniz_equiv _]%auth_both_valid. + Qed. + + Lemma update_value γ_n (n1 n2 m : val) : + own γ_n (● Excl' n1) -∗ own γ_n (◯ Excl' n2) ==∗ own γ_n (● Excl' m) ∗ own γ_n (◯ Excl' m). + Proof. + iIntros "H● H◯". iCombine "H●" "H◯" as "H". rewrite -own_op. iApply (own_update with "H"). + by apply auth_update, option_local_update, exclusive_local_update. + Qed. + + Definition rdcss_content (γ_n : gname) (n : val) := (own γ_n (◯ Excl' n))%I. + + (** Definition of the invariant *) + + (** Extract the TID of the winner from the prophecy value. The winner is the first thread that performs a CAS. *) + Fixpoint proph_extract_winner (pvs : list (val * val)) : option proph_id := + match pvs with + | (_, LitV (LitProphecy tid)) :: _ => Some tid + | _ => None + end. + + Inductive abstract_state : Set := + | Quiescent : val → abstract_state + | Updating : loc → loc → val → val → val → proph_id → abstract_state. + + Definition state_to_val (s : abstract_state) : val := + match s with + | Quiescent n => InjLV n + | Updating l_descr l_m m1 n1 n2 p => InjRV #l_descr + end. + + Definition own_token γ := (own γ (Excl ()))%I. + + Definition pending_state P (n1 : val) (proph_winner : option proph_id) tid_ghost_winner (γ_n γ_a: gname) := + (P ∗ + ⌜match proph_winner with None => True | Some p => p = tid_ghost_winner end⌝ ∗ + own γ_n (● Excl' n1) ∗ + own_token γ_a)%I. + + (* After the prophecy said we are going to win the race, we commit and run the AU, + switching from [pending] to [accepted]. *) + Definition accepted_state Q (proph_winner : option proph_id) (tid_ghost_winner : proph_id) := + ((∃ vs, proph tid_ghost_winner vs) ∗ + Q ∗ + ⌜match proph_winner with None => True | Some tid => tid = tid_ghost_winner end⌝)%I. + + (* The same thread then wins the CmpXchg and moves from [accepted] to [done]. + Then, the [γ_t] token guards the transition to take out [Q]. + Remember that the thread winning the CmpXchg might be just helping. The token + is owned by the thread whose request this is. + In this state, [tid_ghost_winner] serves as a token to make sure that + only the CmpXchg winner can transition to here, and owning half of [l_descr] serves as a + "location" token to ensure there is no ABA going on. Notice how [rdcss_inv] + owns *more than* half of its [l_descr] in the Updating state, + which means we know that the [l_descr] there and here cannot be the same. *) + Definition done_state Qn (l_descr : loc) (tid_ghost_winner : proph_id) (γ_t γ_a: gname) := + ((Qn ∨ own_token γ_t) ∗ (∃ vs, proph tid_ghost_winner vs) ∗ (l_descr ↦{1/2} -) ∗ own_token γ_a)%I. + + (* Invariant expressing the descriptor protocol. + - We always need the [proph] in here so that failing threads coming late can + always resolve their stuff. + - We need a way for anyone who has observed the [done] state to + prove that we will always remain [done]; that's what the one-shot token [γ_s] is for. + - [γ_a] is a token which is owned by the invariant in [pending] and [done] but not in [accepted]. + This permits the CmpXchg winner to gain ownership of the token when moving to [accepted] and + hence ensures that no other thread can move from [accepted] to [done]. + Side remark: One could get rid of this token if one supported fractional ownership of + prophecy resources by only keeping half permission to the prophecy resource + in the invariant in [accepted] while the other half would be kept by the CmpXchg winner. + *) + Definition descr_inv P Q (p : proph_id) n (l_n l_descr : loc) (tid_ghost_winner : proph_id) γ_n γ_t γ_s γ_a: iProp Σ := + (∃ vs, proph p vs ∗ + (own γ_s (Cinl $ Excl ()) ∗ + (l_n ↦{1/2} InjRV #l_descr ∗ ( pending_state P n (proph_extract_winner vs) tid_ghost_winner γ_n γ_a + ∨ accepted_state (Q n) (proph_extract_winner vs) tid_ghost_winner )) + ∨ own γ_s (Cinr $ to_agree ()) ∗ done_state (Q n) l_descr tid_ghost_winner γ_t γ_a))%I. + + Local Hint Extern 0 (environments.envs_entails _ (descr_inv _ _ _ _ _ _ _ _ _ _ _)) => unfold descr_inv. + + Definition pau P Q γ l_m m1 n1 n2 := + (▷ P -∗ ◇ AU << ∀ (m n : val), (gc_mapsto l_m m) ∗ rdcss_content γ n >> @ (⊤∖↑N)∖↑gcN, ∅ + << (gc_mapsto l_m m) ∗ (rdcss_content γ (if (decide ((m = m1) ∧ (n = n1))) then n2 else n)), + COMM Q n >>)%I. + + Definition rdcss_inv γ_n l_n := + (∃ (s : abstract_state), + l_n ↦{1/2} (state_to_val s) ∗ + match s with + | Quiescent n => + (* (InjLV #n) = state_to_val (Quiescent n) *) + (* In this state the CmpXchg which expects to read (InjRV _) in + [complete] fails always.*) + l_n ↦{1/2} (InjLV n) ∗ own γ_n (● Excl' n) + | Updating l_descr l_m m1 n1 n2 p => + ∃ q P Q tid_ghost_winner γ_t γ_s γ_a, + (* (InjRV #l_descr) = state_to_val (Updating l_descr l_m m1 n1 n2 p) *) + (* There are three pieces of per-[descr]-protocol ghost state: + - [γ_t] is a token owned by whoever created this protocol and used + to get out the [Q] in the end. + - [γ_s] reflects whether the protocol is [done] yet or not. + - [γ_a] is a token which is used to ensure that only the CmpXchg winner + can move from the [accepted] to the [done] state. *) + (* We own *more than* half of [l_descr], which shows that this cannot + be the [l_descr] of any [descr] protocol in the [done] state. *) + ⌜val_is_unboxed m1⌝ ∗ + l_descr ↦{1/2 + q} (#l_m, m1, n1, n2, #p)%V ∗ + inv descrN (descr_inv P Q p n1 l_n l_descr tid_ghost_winner γ_n γ_t γ_s γ_a) ∗ + □ pau P Q γ_n l_m m1 n1 n2 ∗ is_gc_loc l_m + end)%I. + + Local Hint Extern 0 (environments.envs_entails _ (rdcss_inv _ _)) => unfold rdcss_inv. + + Definition is_rdcss (γ_n : gname) (l_n: loc) := + (inv rdcssN (rdcss_inv γ_n l_n) ∧ gc_inv ∧ ⌜N ## gcN⌝)%I. + + Global Instance is_rdcss_persistent γ_n l_n: Persistent (is_rdcss γ_n l_n) := _. + + Global Instance rdcss_content_timeless γ_n n : Timeless (rdcss_content γ_n n) := _. + + Global Instance abstract_state_inhabited: Inhabited abstract_state := populate (Quiescent #0). + + Lemma rdcss_content_exclusive γ_n l_n_1 l_n_2 : + rdcss_content γ_n l_n_1 -∗ rdcss_content γ_n l_n_2 -∗ False. + Proof. + iIntros "Hn1 Hn2". iDestruct (own_valid_2 with "Hn1 Hn2") as %?. + done. + Qed. + + (** A few more helper lemmas that will come up later *) + + Lemma mapsto_valid_3 l v1 v2 q : + l ↦ v1 -∗ l ↦{q} v2 -∗ ⌜False⌝. + Proof. + iIntros "Hl1 Hl2". iDestruct (mapsto_valid_2 with "Hl1 Hl2") as %Hv. + apply (iffLR (frac_valid' _)) in Hv. by apply Qp_not_plus_q_ge_1 in Hv. + Qed. + + (** Once a [descr] protocol is [done] (as reflected by the [γ_s] token here), + we can at any later point in time extract the [Q]. *) + Lemma state_done_extract_Q P Q p n l_n l_descr tid_ghost γ_n γ_t γ_s γ_a : + inv descrN (descr_inv P Q p n l_n l_descr tid_ghost γ_n γ_t γ_s γ_a) -∗ + own γ_s (Cinr (to_agree ())) -∗ + □(own_token γ_t ={⊤}=∗ ▷ (Q n)). + Proof. + iIntros "#Hinv #Hs !# Ht". + iInv descrN as (vs) "(Hp & [NotDone | Done])". + * (* Moved back to NotDone: contradiction. *) + iDestruct "NotDone" as "(>Hs' & _ & _)". + iDestruct (own_valid_2 with "Hs Hs'") as %?. contradiction. + * iDestruct "Done" as "(_ & QT & Hrest)". + iDestruct "QT" as "[Qn | >T]"; last first. + { iDestruct (own_valid_2 with "Ht T") as %Contra. + by inversion Contra. } + iSplitR "Qn"; last done. iIntros "!> !>". unfold descr_inv. + iExists _. iFrame "Hp". iRight. + unfold done_state. iFrame "#∗". + Qed. + + (** ** Proof of [complete] *) + + (** The part of [complete] for the succeeding thread that moves from [accepted] to [done] state *) + Lemma complete_succeeding_thread_pending (γ_n γ_t γ_s γ_a: gname) l_n P Q p + (n1 n : val) (l_descr : loc) (tid_ghost : proph_id) Φ : + inv rdcssN (rdcss_inv γ_n l_n) -∗ + inv descrN (descr_inv P Q p n1 l_n l_descr tid_ghost γ_n γ_t γ_s γ_a) -∗ + own_token γ_a -∗ + (□(own_token γ_t ={⊤}=∗ ▷ (Q n1)) -∗ Φ #()) -∗ + own γ_n (● Excl' n) -∗ + WP Resolve (CmpXchg #l_n (InjRV #l_descr) (InjLV n)) #p #tid_ghost ;; #() {{ v, Φ v }}. + Proof. + iIntros "#InvC #InvS Token_a HQ Hn●". wp_bind (Resolve _ _ _)%E. + iInv rdcssN as (s) "(>Hln & Hrest)". + iInv descrN as (vs) "(>Hp & [NotDone | Done])"; last first. + { (* We cannot be [done] yet, as we own the [γ_a] token that serves + as token for that transition. *) + iDestruct "Done" as "(_ & _ & _ & _ & >Token_a')". + by iDestruct (own_valid_2 with "Token_a Token_a'") as %?. + } + iDestruct "NotDone" as "(>Hs & >Hln' & [Pending | Accepted])". + { (* We also cannot be [Pending] any more because we own the [γ_a] token. *) + iDestruct "Pending" as "[_ >(_ & _ & Token_a')]". + by iDestruct (own_valid_2 with "Token_a Token_a'") as %?. + } + (* So, we are [Accepted]. Now we can show that (InjRV l_descr) = (state_to_val s), because + while a [descr] protocol is not [done], it owns enough of + the [rdcss] protocol to ensure that does not move anywhere else. *) + destruct s as [n' | l_descr' l_m' m1' n1' n2' p']. + { simpl. iDestruct (mapsto_agree with "Hln Hln'") as %Heq. inversion Heq. } + iDestruct (mapsto_agree with "Hln Hln'") as %[= ->]. + simpl. + iDestruct "Hrest" as (q P' Q' tid_ghost' γ_t' γ_s' γ_a') "(_ & [>Hld >Hld'] & Hrest)". + (* We perform the CmpXchg. *) + iCombine "Hln Hln'" as "Hln". + wp_apply (wp_resolve with "Hp"); first done. wp_cmpxchg_suc. + iIntros (vs'' ->) "Hp'". simpl. + (* Update to Done. *) + iDestruct "Accepted" as "[Hp_phost_inv [Q Heq]]". + iMod (own_update with "Hs") as "Hs". + { apply (cmra_update_exclusive (Cinr (to_agree ()))). done. } + iDestruct "Hs" as "#Hs'". iModIntro. + iSplitL "Hp_phost_inv Token_a Q Hp' Hld". + (* Update state to Done. *) + { eauto 12 with iFrame. } + iModIntro. iSplitR "HQ". + { iNext. iDestruct "Hln" as "[Hln1 Hln2]". + iExists (Quiescent n). iFrame. } + iApply wp_fupd. wp_seq. iApply "HQ". + iApply state_done_extract_Q; done. + Qed. + + (** The part of [complete] for the failing thread *) + Lemma complete_failing_thread γ_n γ_t γ_s γ_a l_n l_descr P Q p n1 n tid_ghost_inv tid_ghost Φ : + tid_ghost_inv ≠ tid_ghost → + inv rdcssN (rdcss_inv γ_n l_n) -∗ + inv descrN (descr_inv P Q p n1 l_n l_descr tid_ghost_inv γ_n γ_t γ_s γ_a) -∗ + (□(own_token γ_t ={⊤}=∗ ▷ (Q n1)) -∗ Φ #()) -∗ + WP Resolve (CmpXchg #l_n (InjRV #l_descr) (InjLV n)) #p #tid_ghost ;; #() {{ v, Φ v }}. + Proof. + iIntros (Hnl) "#InvC #InvS HQ". wp_bind (Resolve _ _ _)%E. + iInv rdcssN as (s) "(>Hln & Hrest)". + iInv descrN as (vs) "(>Hp & [NotDone | [#Hs Done]])". + { (* [descr] protocol is not done yet: we can show that it + is the active protocol still (l = l'). But then the CmpXchg would + succeed, and our prophecy would have told us that. + So here we can prove that the prophecy was wrong. *) + iDestruct "NotDone" as "(_ & >Hln' & State)". + iDestruct (mapsto_agree with "Hln Hln'") as %[=->]. + iCombine "Hln Hln'" as "Hln". + wp_apply (wp_resolve with "Hp"); first done; wp_cmpxchg_suc. + iIntros (vs'' ->). simpl. + iDestruct "State" as "[Pending | Accepted]". + + iDestruct "Pending" as "[_ [Hvs _]]". iDestruct "Hvs" as %Hvs. by inversion Hvs. + + iDestruct "Accepted" as "[_ [_ Hvs]]". iDestruct "Hvs" as %Hvs. by inversion Hvs. + } + (* So, we know our protocol is [Done]. *) + (* It must be that (state_to_val s) ≠ (InjRV l_descr) because we are in the failing thread. *) + destruct s as [n' | l_descr' l_m' m1' n1' n2' p']. + - (* (injL n) is the current value, hence the CmpXchg fails *) + (* FIXME: proof duplication *) + wp_apply (wp_resolve with "Hp"); first done. wp_cmpxchg_fail. + iIntros (vs'' ->) "Hp". iModIntro. + iSplitL "Done Hp". { by eauto 12 with iFrame. } + iModIntro. + iSplitL "Hln Hrest". { by eauto 12 with iFrame. } + wp_seq. iApply "HQ". + iApply state_done_extract_Q; done. + - (* (injR l_descr') is the current value *) + destruct (decide (l_descr' = l_descr)) as [->|Hn]. + + (* The [descr] protocol is [done] while still being the active protocol + of the [rdcss] instance? Impossible, now we will own more than the whole descriptor location! *) + iDestruct "Done" as "(_ & _ & >Hld & _)". + iDestruct "Hld" as (v') "Hld". + iDestruct "Hrest" as (q P' Q' tid_ghost' γ_t' γ_s' γ_a') "(_ & >[Hld' Hld''] & Hrest)". + iDestruct (mapsto_combine with "Hld Hld'") as "[Hld _]". + rewrite Qp_half_half. iDestruct (mapsto_valid_3 with "Hld Hld''") as "[]". + + (* l_descr' ≠ l_descr: The CmpXchg fails. *) + wp_apply (wp_resolve with "Hp"); first done. wp_cmpxchg_fail. + iIntros (vs'' ->) "Hp". iModIntro. + iSplitL "Done Hp". { by eauto 12 with iFrame. } + iModIntro. + iSplitL "Hln Hrest". { by eauto 12 with iFrame. } + wp_seq. iApply "HQ". + iApply state_done_extract_Q; done. + Qed. + + (** ** Proof of [complete] *) + (* The postcondition basically says that *if* you were the thread to own + this request, then you get [Q]. But we also try to complete other + thread's requests, which is why we cannot ask for the token + as a precondition. *) + Lemma complete_spec (l_n l_m l_descr : loc) (m1 n1 n2 : val) (p : proph_id) γ_n γ_t γ_s γ_a tid_ghost_inv P Q q: + val_is_unboxed m1 → + N ## gcN → + inv rdcssN (rdcss_inv γ_n l_n) -∗ + inv descrN (descr_inv P Q p n1 l_n l_descr tid_ghost_inv γ_n γ_t γ_s γ_a) -∗ + □ pau P Q γ_n l_m m1 n1 n2 -∗ + is_gc_loc l_m -∗ + gc_inv -∗ + {{{ l_descr ↦{q} (#l_m, m1, n1, n2, #p) }}} + complete #l_descr #l_n + {{{ RET #(); □ (own_token γ_t ={⊤}=∗ ▷(Q n1)) }}}. + Proof. + iIntros (Hm_unbox Hdisj) "#InvC #InvS #PAU #isGC #InvGC". + iModIntro. iIntros (Φ) "Hld HQ". wp_lam. wp_let. + wp_bind (! _)%E. wp_load. iClear "Hld". wp_pures. + wp_apply wp_new_proph; first done. + iIntros (vs_ghost tid_ghost) "Htid_ghost". wp_pures. + wp_bind (! _)%E. + (* open outer invariant *) + iInv rdcssN as (s) "(>Hln & Hrest)"=>//. + (* two different proofs depending on whether we are succeeding thread *) + destruct (decide (tid_ghost_inv = tid_ghost)) as [-> | Hnl]. + - (* we are the succeeding thread *) + (* we need to move from [pending] to [accepted]. *) + iInv descrN as (vs) "(>Hp & [(>Hs & >Hln' & [Pending | Accepted]) | [#Hs Done]])". + + (* Pending: update to accepted *) + iDestruct "Pending" as "[P >(Hvs & Hn● & Token_a)]". + iDestruct ("PAU" with "P") as ">AU". + iMod (gc_access with "InvGC") as "Hgc"; first solve_ndisj. + (* open and *COMMIT* AU, sync B location l_n and A location l_m *) + iMod "AU" as (m' n') "[CC [_ Hclose]]". + iDestruct "CC" as "[Hgc_lm Hn◯]". + (* sync B location and update it if required *) + iDestruct (sync_values with "Hn● Hn◯") as %->. + iMod (update_value _ _ _ (if decide (m' = m1 ∧ n' = n') then n2 else n') with "Hn● Hn◯") + as "[Hn● Hn◯]". + (* get access to A location *) + iDestruct ("Hgc" with "Hgc_lm") as "[Hl Hgc_close]". + (* read A location *) + wp_load. + (* sync A location *) + iMod ("Hgc_close" with "Hl") as "[Hgc_lm Hgc_close]". + (* give back access to A location *) + iMod ("Hclose" with "[Hn◯ Hgc_lm]") as "Q"; first by iFrame. + iModIntro. iMod "Hgc_close" as "_". + (* close descr inv *) + iModIntro. iSplitL "Q Htid_ghost Hp Hvs Hs Hln'". + { iModIntro. iNext. iExists _. iFrame "Hp". eauto 12 with iFrame. } + (* close outer inv *) + iModIntro. iSplitR "Token_a HQ Hn●". + { by eauto 12 with iFrame. } + iModIntro. + destruct (decide (m' = m1)) as [-> | ?]; + wp_op; + case_bool_decide; simplify_eq; wp_if; wp_pures; + [rewrite decide_True; last done | rewrite decide_False; last tauto]; + iApply (complete_succeeding_thread_pending + with "InvC InvS Token_a HQ Hn●"). + + (* Accepted: contradiction *) + iDestruct "Accepted" as "[>Htid_ghost_inv _]". + iDestruct "Htid_ghost_inv" as (p') "Htid_ghost_inv". + by iDestruct (proph_exclusive with "Htid_ghost Htid_ghost_inv") as %?. + + (* Done: contradiction *) + iDestruct "Done" as "[QT >[Htid_ghost_inv _]]". + iDestruct "Htid_ghost_inv" as (p') "Htid_ghost_inv". + by iDestruct (proph_exclusive with "Htid_ghost Htid_ghost_inv") as %?. + - (* we are the failing thread *) + (* close invariant *) + iMod (is_gc_access with "InvGC isGC") as (v) "[Hlm Hclose]"; first solve_ndisj. + wp_load. + iMod ("Hclose" with "Hlm") as "_". iModIntro. + iModIntro. + iSplitL "Hln Hrest". + { eauto with iFrame. } + (* two equal proofs depending on value of m1 *) + wp_op. + destruct (decide (v = m1)) as [-> | ]; + case_bool_decide; simplify_eq; wp_if; wp_pures; + iApply (complete_failing_thread + with "InvC InvS HQ"); done. + Qed. + + (** ** Proof of [rdcss] *) + Lemma rdcss_spec γ_n (l_n l_m: loc) (m1 n1 n2: val) : + val_is_unboxed m1 → + val_is_unboxed (InjLV n1) → + is_rdcss γ_n l_n -∗ + <<< ∀ (m n: val), gc_mapsto l_m m ∗ rdcss_content γ_n n >>> + rdcss #l_m #l_n m1 n1 n2 @((⊤∖↑N)∖↑gcN) + <<< gc_mapsto l_m m ∗ rdcss_content γ_n (if decide (m = m1 ∧ n = n1) then n2 else n), RET n >>>. + Proof. + iIntros (Hm1_unbox Hn1_unbox) "Hrdcss". iDestruct "Hrdcss" as "(#InvR & #InvGC & Hdisj)". + iDestruct "Hdisj" as %Hdisj. iIntros (Φ) "AU". + (* allocate fresh descriptor *) + wp_lam. wp_pures. + wp_apply wp_new_proph; first done. + iIntros (proph_values p) "Hp". + wp_let. wp_alloc l_descr as "Hld". + wp_pures. + (* invoke inner recursive function [rdcss_inner] *) + iLöb as "IH". + wp_bind (CmpXchg _ _ _)%E. + (* open outer invariant for the CmpXchg *) + iInv rdcssN as (s) "(>Hln & Hrest)". + destruct s as [n | l_descr' l_m' m1' n1' n2' p']. + - (* l_n ↦ injL n *) + (* a non-value descriptor n is currently stored at l_n *) + iDestruct "Hrest" as ">[Hln' Hn●]". + destruct (decide (n1 = n)) as [-> | Hneq]. + + (* values match -> CmpXchg is successful *) + iCombine "Hln Hln'" as "Hln". + wp_cmpxchg_suc. + (* Take a "peek" at [AU] and abort immediately to get [gc_is_gc f]. *) + iMod "AU" as (b' n') "[[Hf CC] [Hclose _]]". + iDestruct (gc_is_gc with "Hf") as "#Hgc". + iMod ("Hclose" with "[Hf CC]") as "AU"; first by iFrame. + (* Initialize new [descr] protocol .*) + iDestruct (laterable with "AU") as (AU_later) "[AU #AU_back]". + iMod (own_alloc (Excl ())) as (γ_t) "Token_t"; first done. + iMod (own_alloc (Excl ())) as (γ_a) "Token_a"; first done. + iMod (own_alloc (Cinl $ Excl ())) as (γ_s) "Hs"; first done. + iDestruct "Hln" as "[Hln Hln']". + set (winner := default p (proph_extract_winner proph_values)). + iMod (inv_alloc descrN _ (descr_inv AU_later _ _ _ _ _ winner _ _ _ _) + with "[AU Hs Hp Hln' Hn● Token_a]") as "#Hinv". + { + iNext. iExists _. iFrame "Hp". iLeft. iFrame. iLeft. + iFrame. destruct (proph_extract_winner proph_values); simpl; done. + } + iModIntro. iDestruct "Hld" as "[Hld1 [Hld2 Hld3]]". iSplitR "Hld2 Token_t". + { (* close outer invariant *) + iNext. iCombine "Hld1 Hld3" as "Hld1". iExists (Updating l_descr l_m m1 n n2 p). + eauto 15 with iFrame. + } + wp_pures. + wp_apply (complete_spec with "[] [] [] [] [] [$Hld2]");[ done..|]. + iIntros "Ht". iMod ("Ht" with "Token_t") as "Φ". by wp_seq. + + (* values do not match -> CmpXchg fails + we can commit here *) + wp_cmpxchg_fail. + iMod "AU" as (m'' n'') "[[Hm◯ Hn◯] [_ Hclose]]"; simpl. + (* synchronize B location *) + iDestruct (sync_values with "Hn● Hn◯") as %->. + iMod ("Hclose" with "[Hm◯ Hn◯]") as "HΦ". + { destruct (decide _) as [[_ ?] | _]; [done | iFrame ]. } + iModIntro. iSplitR "HΦ". + { iModIntro. iExists (Quiescent n''). iFrame. } + wp_pures. iFrame. + - (* l_n ↦ injR l_ndescr' *) + (* a descriptor l_descr' is currently stored at l_n -> CmpXchg fails + try to help the on-going operation *) + wp_cmpxchg_fail. + iModIntro. + (* extract descr invariant *) + iDestruct "Hrest" as (q P Q tid_ghost γ_t γ_s γ_a) "(#Hm1'_unbox & [Hld1 [Hld2 Hld3]] & #InvS & #P_AU & #P_GC)". + iDestruct "Hm1'_unbox" as %Hm1'_unbox. + iSplitR "AU Hld2 Hld Hp". + (* close invariant, retain some permission to l_descr', so it can be read later *) + { iModIntro. iExists (Updating l_descr' l_m' m1' n1' n2' p'). eauto 15 with iFrame. } + wp_pures. + wp_apply (complete_spec with "[] [] [] [] [] [$Hld2]"); [done..|]. + iIntros "_". wp_seq. wp_pures. + iApply ("IH" with "AU Hp Hld"). + Qed. + + (** ** Proof of [new_rdcss] *) + Lemma new_rdcss_spec (n: val) : + N ## gcN → gc_inv -∗ + {{{ True }}} + new_rdcss n + {{{ l_n γ_n, RET #l_n ; is_rdcss γ_n l_n ∗ rdcss_content γ_n n }}}. + Proof. + iIntros (Hdisj) "#InvGC". iModIntro. + iIntros (Φ) "_ HΦ". wp_lam. wp_apply wp_fupd. + wp_alloc l_n as "Hln". + iMod (own_alloc (● Excl' n ⋅ ◯ Excl' n)) as (γ_n) "[Hn● Hn◯]". + { by apply auth_both_valid. } + iMod (inv_alloc rdcssN _ (rdcss_inv γ_n l_n) + with "[Hln Hn●]") as "#InvR". + { iNext. iDestruct "Hln" as "[Hln1 Hln2]". + iExists (Quiescent n). iFrame. } + iModIntro. + iApply ("HΦ" $! l_n γ_n). + iSplitR; last by iFrame. by iFrame "#". + Qed. + + (** ** Proof of [get] *) + Lemma get_spec γ_n l_n : + is_rdcss γ_n l_n -∗ + <<< ∀ (n : val), rdcss_content γ_n n >>> + get #l_n @(⊤∖↑N) + <<< rdcss_content γ_n n, RET n >>>. + Proof. + iIntros "Hrdcss". iDestruct "Hrdcss" as "(#InvR & #InvGC & Hdisj)". + iDestruct "Hdisj" as %Hdisj. iIntros (Φ) "AU". + iLöb as "IH". wp_lam. repeat (wp_proj; wp_let). wp_bind (! _)%E. + iInv rdcssN as (s) "(>Hln & Hrest)". + wp_load. + destruct s as [n | l_descr l_m m1 n1 n2 p]. + - iMod "AU" as (au_n) "[Hn◯ [_ Hclose]]"; simpl. + iDestruct "Hrest" as "[Hln' Hn●]". + iDestruct (sync_values with "Hn● Hn◯") as %->. + iMod ("Hclose" with "Hn◯") as "HΦ". + iModIntro. iSplitR "HΦ". { + iNext. iExists (Quiescent au_n). iFrame. + } + wp_match. iApply "HΦ". + - iDestruct "Hrest" as (q P Q tid_ghost γ_t γ_s γ_a) "(#Hm1_unbox & [Hld [Hld' Hld'']] & #InvS & #PAU & #GC)". + iDestruct "Hm1_unbox" as %Hm1_unbox. + iModIntro. iSplitR "AU Hld'". { + iNext. iExists (Updating l_descr l_m m1 n1 n2 p). eauto 15 with iFrame. + } + wp_match. + wp_apply (complete_spec with "[] [] [] [] [] [$Hld']"); [done..|]. + iIntros "Ht". wp_seq. iApply "IH". iApply "AU". + Qed. + +End rdcss. + +Definition atomic_rdcss `{!heapG Σ, !rdcssG Σ, !gcG Σ} : + spec.atomic_rdcss Σ := + {| spec.new_rdcss_spec := new_rdcss_spec; + spec.rdcss_spec := rdcss_spec; + spec.get_spec := get_spec; + spec.rdcss_content_exclusive := rdcss_content_exclusive |}. + +Typeclasses Opaque rdcss_content is_rdcss. diff --git a/theories/logatom/rdcss/spec.v b/theories/logatom/rdcss/spec.v new file mode 100644 index 0000000000000000000000000000000000000000..10c4a923787f6e03d0c26679ed2f525e7a891c0d --- /dev/null +++ b/theories/logatom/rdcss/spec.v @@ -0,0 +1,49 @@ +From stdpp Require Import namespaces. +From iris.heap_lang Require Export lifting notation. +From iris.program_logic Require Export atomic. +From iris_examples.logatom.lib Require Export gc. +Set Default Proof Using "Type". + +(** A general logically atomic interface for conditional increment. *) +Record atomic_rdcss {Σ} `{!heapG Σ, !gcG Σ} := AtomicRdcss { + (* -- operations -- *) + new_rdcss : val; + rdcss: val; + get : val; + (* -- other data -- *) + name : Type; + name_eqdec : EqDecision name; + name_countable : Countable name; + (* -- predicates -- *) + is_rdcss (N : namespace) (γ : name) (l_n : loc) : iProp Σ; + rdcss_content (γ : name) (n : val) : iProp Σ; + (* -- predicate properties -- *) + is_rdcss_persistent N γ l_n : Persistent (is_rdcss N γ l_n); + rdcss_content_timeless γ n : Timeless (rdcss_content γ n); + rdcss_content_exclusive γ n1 n2 : rdcss_content γ n1 -∗ rdcss_content γ n2 -∗ False; + (* -- operation specs -- *) + new_rdcss_spec N (n : val): + N ## gcN → gc_inv -∗ + {{{ True }}} + new_rdcss n + {{{ l_n γ, RET #l_n ; is_rdcss N γ l_n ∗ rdcss_content γ n }}}; + rdcss_spec N γ (l_n l_m : loc) (m1 n1 n2 : val): + val_is_unboxed m1 → + val_is_unboxed (InjLV n1) → + is_rdcss N γ l_n -∗ + <<< ∀ (m n: val), gc_mapsto l_m m ∗ rdcss_content γ n >>> + rdcss #l_m #l_n m1 n1 n2 @((⊤∖↑N)∖↑gcN) + <<< gc_mapsto l_m m ∗ rdcss_content γ (if decide (m = m1 ∧ n = n1) then n2 else n), RET n >>>; + get_spec N γ (l_n : loc): + is_rdcss N γ l_n -∗ + <<< ∀ (n : val), rdcss_content γ n >>> + get #l_n @(⊤∖↑N) + <<< rdcss_content γ n, RET n >>>; +}. +Arguments atomic_rdcss _ {_} {_}. + + +Existing Instances + is_rdcss_persistent rdcss_content_timeless + name_countable name_eqdec. + diff --git a/theories/logatom/snapshot/atomic_snapshot.v b/theories/logatom/snapshot/atomic_snapshot.v index 026750bd8f2e0cf4c4dced3997234ff2bc65a262..391caf1c59436ced42effafeaf13d273d3564693 100644 --- a/theories/logatom/snapshot/atomic_snapshot.v +++ b/theories/logatom/snapshot/atomic_snapshot.v @@ -9,18 +9,78 @@ From iris_examples.logatom.snapshot Require Import spec. Set Default Proof Using "Type". (** Specifying snapshots with histories - Implementing atomic pair snapshot data structure from Sergey et al. (ESOP 2015) *) + Inspired by atomic pair snapshot data structure from Sergey et al. (ESOP 2015) *) + + +(* + new_snapshot v := + ref (ref (v, 0)) + *) +Definition new_snapshot : val := + λ: "v", + ref (ref ("v", #0)). + +(* + write xp x := + let xp1 = !xp in + let v = (!xp1).2 + let xp2 = ref (x, v + 1) + // This is where we need the nested references: the CAS compares the old + // inner reference with the new one, because it cannot atomically compared + // a pair of values. + if CAS xp xp1 xp2 + then () + else writeX (xp, yp) x + *) +Definition write : val := + rec: "write" "xp" "x" := + let: "xp1" := !"xp" in + let: "v" := Snd (!"xp1") in + let: "xp2" := ref ("x", "v" + #1) in + if: CAS "xp" "xp1" "xp2" + then #() + else "write" "xp" "x". + +(* + read xp := + (!!xp).1 + *) +Definition read : val := + λ: "xp", + Fst !(!"xp"). + +(* + read_with xp l := + let (x, v) = !!xp in + let y = !l in + let (_, v') = !!xp in + if v = v' + then (x, y) + else read_with l + *) +Definition read_with_proph : val := + rec: "read_with" "xp" "l" := + let: "proph" := NewProph in + let: "x" := ! !"xp" in + let: "y" := !"l" in + let: "x'" := ! !"xp" in + let: "v_eq" := Snd "x" = Snd "x'" in + resolve_proph: "proph" to: "v_eq" ;; + if: "v_eq" + then (Fst "x", "y") + else "read_with" "xp" "l". + (** The CMRA & functor we need. *) -Definition timestampUR := gmapUR Z $ agreeR valC. +Definition timestampUR := gmapUR Z $ agreeR valO. Class atomic_snapshotG Σ := AtomicSnapshotG { - atomic_snapshot_stateG :> inG Σ $ authR $ optionUR $ exclR $ prodC valC valC; + atomic_snapshot_stateG :> inG Σ $ authR $ optionUR $ exclR $ valO; atomic_snapshot_timestampG :> inG Σ $ authR $ timestampUR }. Definition atomic_snapshotΣ : gFunctors := - #[GFunctor (authR $ optionUR $ exclR $ prodC valC valC); GFunctor (authR timestampUR)]. + #[GFunctor (authR $ optionUR $ exclR $ valO); GFunctor (authR timestampUR)]. Instance subG_atomic_snapshotΣ {Σ} : subG atomic_snapshotΣ Σ → atomic_snapshotG Σ. Proof. solve_inG. Qed. @@ -29,152 +89,66 @@ Section atomic_snapshot. Context {Σ} `{!heapG Σ, !atomic_snapshotG Σ}. - (* - newPair x y := - (ref (ref (x, 0)), ref y) - *) - Definition newPair : val := - λ: "args", - let: "x" := Fst "args" in - let: "y" := Snd "args" in - (ref (ref ("x", #0)), ref "y"). - - (* - writeY (xp, yp) y := - yp <- y - *) - Definition writeY : val := - λ: "args", - let: "p" := Fst "args" in - Snd "p" <- Snd "args". - - (* - writeX (xp, yp) x := - let xp1 = !xp in - let v = (!xp1).2 - let xp2 = ref (x, v + 1) - if CAS xp xp1 xp2 - then () - else writeX (xp, yp) x - *) - Definition writeX : val := - rec: "writeX" "args" := - let: "p" := Fst "args" in - let: "x" := Snd "args" in - let: "xp" := Fst "p" in - let: "xp1" := !"xp" in - let: "v" := Snd (!"xp1") in - let: "xp2" := ref ("x", "v" + #1) in - if: CAS "xp" "xp1" "xp2" - then #() - else "writeX" "args". - - (* - readX (xp, yp) := - !!xp - *) - Definition readX : val := - λ: "p", - let: "xp" := Fst "p" in - !(!"xp"). - - Definition readY : val := - λ: "p", - let: "yp" := Snd "p" in - !"yp". - - (* - readPair l := - let (x, v) = readX l in - let y = readY l in - let (_, v') = readX l in - if v = v' - then (x, y) - else readPair l - *) - Definition readPair : val := - rec: "readPair" "l" := - let: "x" := readX "l" in - let: "y" := readY "l" in - let: "x'" := readX "l" in - if: Snd "x" = Snd "x'" - then (Fst "x", Fst "y") - else "readPair" "l". - - Definition readPair_proph : val := - rec: "readPair" "l" := - let: "xv1" := readX "l" in - let: "proph" := NewProph in - let: "y" := readY "l" in - let: "xv2" := readX "l" in - let: "v2" := Snd "xv2" in - let: "v_eq" := Snd "xv1" = "v2" in - resolve_proph: "proph" to: "v_eq" ;; - if: "v_eq" - then (Fst "xv1", "y") - else "readPair" "l". - Variable N: namespace. Definition gmap_to_UR T : timestampUR := to_agree <$> T. - Definition pair_inv γ1 γ2 l1 l2 : iProp Σ := - (∃ q (l1':loc) (T : gmap Z val) x y (t : Z), + Definition snapshot_inv γ1 γ2 l1 : iProp Σ := + (∃ q (l1':loc) (T : gmap Z val) x (t : Z), (* we add the q to make the l1' map fractional. that way, we can take a fraction of the l1' map out of the invariant and do a case analysis on whether the pointer is the same throughout invariant openings *) - l1 ↦ #l1' ∗ l1' ↦{q} (x, #t) ∗ l2 ↦ y ∗ - own γ1 (● Excl' (x, y)) ∗ own γ2 (● gmap_to_UR T) ∗ + l1 ↦ #l1' ∗ l1' ↦{q} (x, #t) ∗ + own γ1 (● Excl' x) ∗ own γ2 (● gmap_to_UR T) ∗ ⌜T !! t = Some x⌝ ∗ ⌜forall (t' : Z), t' ∈ dom (gset Z) T → (t' <= t)%Z⌝)%I. - Definition is_pair (γs: gname * gname) (p : val) := - (∃ (l1 l2 : loc), ⌜p = (#l1, #l2)%V⌝ ∗ inv N (pair_inv γs.1 γs.2 l1 l2))%I. + Definition is_snapshot (γs: gname * gname) (p : val) := + (∃ (l1 : loc), ⌜p = #l1%V⌝ ∗ inv N (snapshot_inv γs.1 γs.2 l1))%I. - Global Instance is_pair_persistent γs p : Persistent (is_pair γs p) := _. + Global Instance is_snapshot_persistent γs p : Persistent (is_snapshot γs p) := _. - Definition pair_content (γs : gname * gname) (a : val * val) := + Definition snapshot_content (γs : gname * gname) (a : val) := (own γs.1 (◯ Excl' a))%I. - Global Instance pair_content_timeless γs a : Timeless (pair_content γs a) := _. + Global Instance snapshot_content_timeless γs a : Timeless (snapshot_content γs a) := _. - Lemma pair_content_exclusive γs a1 a2 : - pair_content γs a1 -∗ pair_content γs a2 -∗ False. + Lemma snapshot_content_exclusive γs a1 a2 : + snapshot_content γs a1 -∗ snapshot_content γs a2 -∗ False. Proof. iIntros "H1 H2". iDestruct (own_valid_2 with "H1 H2") as %[]. Qed. Definition new_timestamp t v : gmap Z val := {[ t := v ]}. - Lemma newPair_spec (v1 v2 : val) : + Lemma new_snapshot_spec (v : val) : {{{ True }}} - newPair (v1, v2)%V - {{{ γs p, RET p; is_pair γs p ∗ pair_content γs (v1, v2) }}}. + new_snapshot v + {{{ γs p, RET p; is_snapshot γs p ∗ snapshot_content γs v }}}. Proof. - iIntros (Φ _) "Hp". rewrite /newPair. wp_lam. + iIntros (Φ _) "Hx". rewrite /new_snapshot. wp_lam. repeat (wp_proj; wp_let). iApply wp_fupd. wp_alloc lx' as "Hlx'". wp_alloc lx as "Hlx". - wp_alloc ly as "Hly". - set (Excl' (v1, v2)) as p. - iMod (own_alloc (● p ⋅ ◯ p)) as (γ1) "[Hp⚫ Hp◯]". { + set (Excl' v) as p. + iMod (own_alloc (● p ⋅ ◯ p)) as (γ1) "[Hx⚫ Hx◯]". { rewrite /p. apply auth_both_valid. split; done. } - set (new_timestamp 0 v1) as t. + set (new_timestamp 0 v) as t. iMod (own_alloc (● gmap_to_UR t ⋅ ◯ gmap_to_UR t)) as (γ2) "[Ht⚫ Ht◯]". { rewrite /t /new_timestamp. apply auth_both_valid. split; first done. rewrite /gmap_to_UR map_fmap_singleton. apply singleton_valid. done. } - wp_pures. iApply ("Hp" $! (γ1, γ2)). - iMod (inv_alloc N _ (pair_inv γ1 γ2 _ _) with "[-Hp◯ Ht◯]") as "#Hinv". { - iNext. rewrite /pair_inv. iExists _, _, _, _, _. iExists 0. iFrame. + wp_pures. iApply ("Hx" $! (γ1, γ2)). + iMod (inv_alloc N _ (snapshot_inv γ1 γ2 _) with "[-Hx◯ Ht◯]") as "#Hinv". { + iNext. rewrite /snapshot_inv. iExists _, _, _, _, 0. iFrame. iPureIntro. split; first done. intros ?. subst t. rewrite /new_timestamp dom_singleton. rewrite elem_of_singleton. lia. } - iModIntro. iSplitR. rewrite /is_pair. repeat (iExists _). iSplitL; eauto. - rewrite /pair_content. rewrite /p. iApply "Hp◯". + iModIntro. rewrite /is_snapshot /snapshot_content. iFrame. iExists _. + iSplitL; first done. done. Qed. Lemma excl_update γ n' n m : @@ -232,75 +206,50 @@ Section atomic_snapshot. by intros ->. Qed. - Lemma writeY_spec γ (y2: val) p : - is_pair γ p -∗ - <<< ∀ x y : val, pair_content γ (x, y) >>> - writeY (p, y2)%V + Lemma write_spec γ (x2: val) p : + is_snapshot γ p -∗ + <<< ∀ x : val, snapshot_content γ x >>> + write p x2 @ ⊤∖↑N - <<< pair_content γ (x, y2), RET #() >>>. + <<< snapshot_content γ x2, RET #() >>>. Proof. - iIntros "Hp". iIntros (Φ) "AU". - iDestruct "Hp" as (l1 l2 ->) "#Hinv". wp_pures. - wp_lam. wp_pures. - iApply wp_fupd. - iInv N as (q l1' T x) "Hinvp". - iDestruct "Hinvp" as (y v') "[Hl1 [Hl1' [Hl2 [Hp⚫ Htime]]]]". - wp_store. - iMod "AU" as (xv yv) "[Hpair [_ Hclose]]". - rewrite /pair_content. - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. - iMod (excl_update _ (x, y2) with "Hp⚫ Hpair") as "[Hp⚫ Hp◯]". - iMod ("Hclose" with "Hp◯") as "HΦ". - iModIntro. iSplitR "HΦ"; last done. - iNext. unfold pair_inv. eauto 7 with iFrame. - Qed. - - Lemma writeX_spec γ (x2: val) p : - is_pair γ p -∗ - <<< ∀ x y : val, pair_content γ (x, y) >>> - writeX (p, x2)%V - @ ⊤∖↑N - <<< pair_content γ (x2, y), RET #() >>>. - Proof. - iIntros "Hp". iIntros (Φ) "AU". iLöb as "IH". - iDestruct "Hp" as (l1 l2 ->) "#Hinv". wp_pures. wp_lam. wp_pures. + iIntros "Hx". iIntros (Φ) "AU". iLöb as "IH". + iDestruct "Hx" as (l1 ->) "#Hinv". wp_pures. wp_lam. wp_pures. (* first read *) (* open invariant *) - wp_bind (!_)%E. iInv N as (q l1' T x) "Hinvp". - iDestruct "Hinvp" as (y v') "[Hl1 [Hl1' [Hl2 [Hp⚫ Htime]]]]". + wp_bind (!_)%E. iInv N as (q l1' T x v') "[Hl1 [Hl1' [Hx⚫ Htime]]]". wp_load. iDestruct "Hl1'" as "[Hl1'frac1 Hl1'frac2]". iModIntro. iSplitR "AU Hl1'frac2". (* close invariant *) - { iNext. rewrite /pair_inv. eauto 10 with iFrame. } + { iNext. rewrite /snapshot_inv. eauto 10 with iFrame. } wp_let. wp_bind (!_)%E. clear T. wp_load. wp_proj. wp_let. wp_op. wp_alloc l1'new as "Hl1'new". wp_let. (* CAS *) - wp_bind (CAS _ _ _)%E. + wp_bind (CmpXchg _ _ _)%E. (* open invariant *) - iInv N as (q' l1'' T x') ">Hinvp". - iDestruct "Hinvp" as (y' v'') "[Hl1 [Hl1'' [Hl2 [Hp⚫ [Ht● Ht]]]]]". + iInv N as (q' l1'' T x' v'') ">[Hl1 [Hl1'' [Hx⚫ [Ht● Ht]]]]". iDestruct "Ht" as %[Ht Hvt]. destruct (decide (l1'' = l1')) as [-> | Hn]. - - wp_cas_suc. + - wp_cmpxchg_suc. iDestruct (mapsto_agree with "Hl1'frac2 Hl1''") as %[= -> ->]. iClear "Hl1'frac2". (* open AU *) - iMod "AU" as (xv yv) "[Hpair [_ Hclose]]". - (* update pair ghost state to (x2, y') *) - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. - iMod (excl_update _ (x2, y') with "Hp⚫ Hpair") as "[Hp⚫ Hp◯]". + iMod "AU" as (xv) "[Hx [_ Hclose]]". + (* update snapshot ghost state to (x2, y') *) + iDestruct (excl_sync with "Hx⚫ Hx") as %[= ->]. + iMod (excl_update _ x2 with "Hx⚫ Hx") as "[Hx⚫ Hx◯]". (* close AU *) - iMod ("Hclose" with "Hp◯") as "HΦ". + iMod ("Hclose" with "Hx◯") as "HΦ". (* update timestamp *) iMod (timestamp_update _ T (v'' + 1)%Z x2 with "[Ht●]") as "Ht". { eapply (not_elem_of_dom (D:=gset Z) T). intros Hd. specialize (Hvt _ Hd). omega. } { done. } (* close invariant *) iModIntro. iSplitR "HΦ". - + iNext. rewrite /pair_inv. + + iNext. rewrite /snapshot_inv. set (<[ v'' + 1 := x2]> T) as T'. - iExists 1%Qp, l1'new, T', x2, y', (v'' + 1). + iExists 1%Qp, l1'new, T', x2, (v'' + 1). iFrame. iPureIntro. split. * apply: lookup_insert. @@ -311,58 +260,34 @@ Section atomic_snapshot. rewrite Hd in Hv. clear Hd. apply elem_of_union in Hv. destruct Hv as [Hv%elem_of_singleton_1 | Hv]; first done. specialize (Hvt _ Hv). lia. - + wp_if. done. - - wp_cas_fail. iModIntro. iSplitR "AU". - + iNext. rewrite /pair_inv. eauto 10 with iFrame. - + wp_if. iApply "IH"; last eauto. rewrite /is_pair. eauto. - Qed. - - Lemma readY_spec γ p : - is_pair γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - readY p - @ ⊤∖↑N - <<< pair_content γ (v1, v2), RET v2 >>>. - Proof. - iIntros "Hp". iIntros (Φ) "AU". - iDestruct "Hp" as (l1 l2 ->) "#Hinv". - repeat (wp_lam; wp_proj). wp_let. - iApply wp_fupd. - iInv N as (q l1' T x) "Hinvp". - iDestruct "Hinvp" as (y v') "[Hl1 [Hl1' [Hl2 [Hp⚫ Htime]]]]". - wp_load. - iMod "AU" as (xv yv) "[Hpair [_ Hclose]]". - rewrite /pair_content. - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. - iMod ("Hclose" with "Hpair") as "HΦ". - iModIntro. iSplitR "HΦ"; last done. - iNext. unfold pair_inv. eauto 7 with iFrame. + + wp_pures. done. + - wp_cmpxchg_fail. iModIntro. iSplitR "AU". + + iNext. rewrite /snapshot_inv. eauto 10 with iFrame. + + wp_pures. iApply "IH"; last eauto. rewrite /is_snapshot. eauto. Qed. - Lemma readX_spec γ p : - is_pair γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - readX p + Lemma read_spec γ p : + is_snapshot γ p -∗ + <<< ∀ v : val, snapshot_content γ v >>> + read p @ ⊤∖↑N - <<< ∃ (t: Z), pair_content γ (v1, v2), RET (v1, #t) >>>. + <<< snapshot_content γ v, RET v >>>. Proof. - iIntros "Hp". iIntros (Φ) "AU". - iDestruct "Hp" as (l1 l2 ->) "#Hinv". - repeat (wp_lam; wp_proj). wp_let. wp_bind (! #l1)%E. + iIntros "Hx". iIntros (Φ) "AU". + iDestruct "Hx" as (l1 ->) "#Hinv". + wp_lam. wp_bind (! #l1)%E. (* open invariant for 1st read *) - iInv N as (q l1' T x) ">Hinvp". - iDestruct "Hinvp" as (y v') "[Hl1 [Hl1' [Hl2 [Hp⚫ Htime]]]]". + iInv N as (q l1' T x v') ">[Hl1 [Hl1' [Hx⚫ Htime]]]". wp_load. iDestruct "Hl1'" as "[Hl1' Hl1'frac]". - iMod "AU" as (xv yv) "[Hpair [_ Hclose]]". - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. - iMod ("Hclose" with "Hpair") as "HΦ". + iMod "AU" as (xv) "[Hx [_ Hclose]]". + iDestruct (excl_sync with "Hx⚫ Hx") as %[= ->]. + iMod ("Hclose" with "Hx") as "HΦ". (* close invariant *) iModIntro. iSplitR "HΦ Hl1'". { - iNext. unfold pair_inv. eauto 7 with iFrame. + iNext. unfold snapshot_inv. eauto 7 with iFrame. } - iApply wp_fupd. clear T y. - wp_load. eauto. + iApply wp_fupd. wp_load. wp_pures. eauto. Qed. Definition prophecy_to_bool (v : list (val * val)) : bool := @@ -371,67 +296,63 @@ Section atomic_snapshot. | _ => false end. - Lemma readPair_spec γ p : - is_pair γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - readPair_proph p + Lemma read_with_spec γ p (l : loc) : + is_snapshot γ p -∗ + <<< ∀ v1 v2 : val, snapshot_content γ v1 ∗ l ↦ v2 >>> + read_with_proph p #l @ ⊤∖↑N - <<< pair_content γ (v1, v2), RET (v1, v2) >>>. + <<< snapshot_content γ v1 ∗ l ↦ v2, RET (v1, v2) >>>. Proof. - iIntros "Hp". iIntros (Φ) "AU". iLöb as "IH". + iIntros "Hx". iIntros (Φ) "AU". iLöb as "IH". wp_lam. wp_pures. + (* ************ new prophecy ********** *) + wp_apply wp_new_proph; first done. + iIntros (proph_val proph) "Hpr". wp_pures. (* ************ 1st readX ********** *) - iDestruct "Hp" as (l1 l2 ->) "#Hinv". repeat (wp_lam; wp_pures). - wp_bind (! #l1)%E. + iDestruct "Hx" as (l1 ->) "#Hinv". wp_bind (! #l1)%E. (* open invariant for 1st read *) - iInv N as (q_x1 l1' T_x x1) ">Hinvp". - iDestruct "Hinvp" as (y_x1 v_x1) "[Hl1 [Hl1' [Hl2 [Hp⚫ [Ht_x Htime_x]]]]]". + iInv N as (q_x1 l1' T_x x1 v_x1) ">[Hl1 [Hl1' [Hx⚫ [Ht_x Htime_x]]]]". iDestruct "Htime_x" as %[Hlookup_x Hdom_x]. wp_load. iDestruct "Hl1'" as "[Hl1' Hl1'frac]". - iMod "AU" as (xv yv) "[Hpair [Hclose _]]". - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. - iMod ("Hclose" with "Hpair") as "AU". + iMod "AU" as (xv yv) "[[Hx Hy] [Hclose _]]". + iDestruct (excl_sync with "Hx⚫ Hx") as %[= ->]. + iMod ("Hclose" with "[$Hx $Hy]") as "AU". (* duplicate timestamp T_x1 *) iMod (timestamp_dupl with "Ht_x") as "[Ht_x1⚫ Ht_x1◯]". (* close invariant *) - iModIntro. iSplitR "AU Hl1' Ht_x1◯". { - iNext. unfold pair_inv. eauto 8 with iFrame. + iModIntro. iSplitR "AU Hl1' Ht_x1◯ Hpr". { + iNext. unfold snapshot_inv. eauto 8 with iFrame. } wp_load. wp_let. - (* ************ new prophecy ********** *) - wp_apply wp_new_proph; first done. - iIntros (proph_val proph) "Hpr". - wp_let. (* ************ readY ********** *) repeat (wp_lam; wp_pures). wp_bind (!_)%E. - iInv N as (q_y l1'_y T_y x_y) ">Hinvp". - iDestruct "Hinvp" as (y_y v_y) "[Hl1 [Hl1'_y [Hl2 [Hp⚫ [Ht_y Htime_y]]]]]". + iInv N as (q_y l1'_y T_y x_y v_y) ">[Hl1 [Hl1'_y [Hx⚫ [Ht_y Htime_y]]]]". iDestruct "Htime_y" as %[Hlookup_y Hdom_y]. - wp_load. (* linearization point *) - iMod "AU" as (xv yv) "[Hpair Hclose]". - rewrite /pair_content. - iDestruct (excl_sync with "Hp⚫ Hpair") as %[= -> ->]. + clear yv. + iMod "AU" as (xv yv) "[[Hx Hy] Hclose]". + wp_load. + rewrite /snapshot_content. + iDestruct (excl_sync with "Hx⚫ Hx") as %[= ->]. destruct (prophecy_to_bool proph_val) eqn:Hproph. - (* prophecy value is predicting that timestamp has not changed, so we commit *) (* committing AU *) - iMod ("Hclose" with "Hpair") as "HΦ". + iMod ("Hclose" with "[$Hx $Hy]") as "HΦ". (* duplicate timestamp T_y *) iMod (timestamp_dupl with "Ht_y") as "[Ht_y● Ht_y◯]". (* show that T_x <= T_y *) iDestruct (timestamp_sub with "[Ht_y● Ht_x1◯]") as "#Hs"; first by iFrame. iDestruct "Hs" as %Hs. - iModIntro. + iModIntro. iModIntro. (* closing invariant *) iSplitR "HΦ Hl1' Ht_x1◯ Ht_y◯ Hpr". - { iNext. unfold pair_inv. eauto 10 with iFrame. } + { iNext. unfold snapshot_inv. eauto 10 with iFrame. } wp_let. (* ************ 2nd readX ********** *) repeat (wp_lam; wp_pures). wp_bind (! #l1)%E. (* open invariant *) - iInv N as (q_x2 l1'_x2 T_x2 x2) ">Hinvp". - iDestruct "Hinvp" as (y_x2 v_x2) "[Hl1 [Hl1'_x2 [Hl2 [Hp⚫ [Ht_x2 Htime_x2]]]]]". + iInv N as (q_x2 l1'_x2 T_x2 x2 v_x2) ">[Hl1 [Hl1'_x2 [Hx⚫ [Ht_x2 Htime_x2]]]]". iDestruct "Htime_x2" as %[Hlookup_x2 Hdom_x2]. iDestruct "Hl1'_x2" as "[Hl1'_x2 Hl1'_x2_frag]". wp_load. @@ -439,10 +360,10 @@ Section atomic_snapshot. iDestruct (timestamp_sub with "[Ht_x2 Ht_y◯]") as "#Hs'"; first by iFrame. iDestruct "Hs'" as %Hs'. iModIntro. iSplitR "HΦ Hl1'_x2_frag Hpr". { - iNext. unfold pair_inv. eauto 8 with iFrame. + iNext. unfold snapshot_inv. eauto 8 with iFrame. } - wp_load. wp_let. wp_proj. wp_let. wp_proj. wp_op. - case_bool_decide; wp_let; wp_apply (wp_resolve_proph with "Hpr"); + wp_load. wp_pures. + case_bool_decide; wp_apply (wp_resolve_proph with "Hpr"); iIntros (vs') "[Eq _]"; iDestruct "Eq" as %->; wp_seq; wp_if. + inversion H; subst; clear H. wp_pures. assert (v_x2 = v_y) as ->. { @@ -461,38 +382,38 @@ Section atomic_snapshot. } done. + inversion Hproph. - - iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "Hpair") as "AU". - iModIntro. + - (* prophecy value is predicting that timestamp has not changed, so we abort *) + iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "[$Hx $Hy]") as "AU". + iModIntro. iModIntro. (* closing invariant *) iSplitR "AU Hpr". - { iNext. unfold pair_inv. eauto 10 with iFrame. } + { iNext. unfold snapshot_inv. eauto 10 with iFrame. } wp_let. (* ************ 2nd readX ********** *) - repeat (wp_lam; wp_proj). wp_let. wp_bind (! #l1)%E. + wp_bind (! #l1)%E. (* open invariant *) - iInv N as (q_x2 l1'_x2 T_x2 x2) ">Hinvp". - iDestruct "Hinvp" as (y_x2 v_x2) "[Hl1 [Hl1'_x2 [Hl2 [Hp⚫ [Ht_x2 Htime_x2]]]]]". + iInv N as (q_x2 l1'_x2 T_x2 x2 v_x2) "[Hl1 [Hl1'_x2 [Hx⚫ [Ht_x2 Htime_x2]]]]". iDestruct "Hl1'_x2" as "[Hl1'_x2 Hl1'_x2_frag]". wp_load. iModIntro. iSplitR "AU Hl1'_x2_frag Hpr". { - iNext. unfold pair_inv. eauto 8 with iFrame. + iNext. unfold snapshot_inv. eauto 8 with iFrame. } - wp_load. repeat (wp_let; wp_proj). wp_op. wp_let. + wp_load. wp_pures. wp_apply (wp_resolve_proph with "Hpr"). iIntros (vs') "[Heq _]"; iDestruct "Heq" as %->. case_bool_decide. + inversion H; subst; clear H. inversion Hproph. - + wp_seq. wp_if. iApply "IH"; rewrite /is_pair; eauto. + + wp_seq. wp_if. iApply "IH"; rewrite /is_snapshot; eauto. Qed. End atomic_snapshot. -Definition atomic_snapshot `{!heapG Σ, atomic_snapshotG Σ} : +Program Definition atomic_snapshot `{!heapG Σ, atomic_snapshotG Σ} : spec.atomic_snapshot Σ := - {| spec.newPair_spec := newPair_spec; - spec.writeX_spec := writeX_spec; - spec.writeY_spec := writeY_spec; - spec.readPair_spec := readPair_spec; - spec.pair_content_exclusive := pair_content_exclusive |}. + {| spec.new_snapshot_spec := new_snapshot_spec; + spec.write_spec := write_spec; + spec.read_spec := read_spec; + spec.read_with_spec := read_with_spec; + spec.snapshot_content_exclusive := snapshot_content_exclusive |}. -Typeclasses Opaque pair_content is_pair. +Typeclasses Opaque snapshot_content is_snapshot. diff --git a/theories/logatom/snapshot/spec.v b/theories/logatom/snapshot/spec.v index 1866b8ab62bfcd0752d6b875c96ed722b2d0d243..10c251b0419a9462f551de357012ea071cbcf337 100644 --- a/theories/logatom/snapshot/spec.v +++ b/theories/logatom/snapshot/spec.v @@ -11,46 +11,46 @@ Set Default Proof Using "Type". Implementing atomic pair snapshot data structure from Sergey et al. (ESOP 2015) *) Record atomic_snapshot {Σ} `{!heapG Σ} := AtomicSnapshot { - newPair : val; - writeX : val; - writeY : val; - readPair : val; + new_snapshot : val; + read : val; + write : val; + read_with : val; (* other data *) name: Type; name_eqdec : EqDecision name; name_countable : Countable name; (* predicates *) - is_pair (N : namespace) (γ : name) (p : val) : iProp Σ; - pair_content (γ : name) (a: val * val) : iProp Σ; + is_snapshot (N : namespace) (γ : name) (p : val) : iProp Σ; + snapshot_content (γ : name) (a: val) : iProp Σ; (* predicate properties *) - is_pair_persistent N γ p : Persistent (is_pair N γ p); - pair_content_timeless γ a : Timeless (pair_content γ a); - pair_content_exclusive γ a1 a2 : - pair_content γ a1 -∗ pair_content γ a2 -∗ False; + is_snapshot_persistent N γ p : Persistent (is_snapshot N γ p); + snapshot_content_timeless γ a : Timeless (snapshot_content γ a); + snapshot_content_exclusive γ a1 a2 : + snapshot_content γ a1 -∗ snapshot_content γ a2 -∗ False; (* specs *) - newPair_spec N (v1 v2 : val) : - {{{ True }}} newPair (v1, v2)%V {{{ γ p, RET p; is_pair N γ p ∗ pair_content γ (v1, v2) }}}; - writeX_spec N γ (v: val) p : - is_pair N γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - writeX (p, v)%V + new_snapshot_spec N (v : val) : + {{{ True }}} new_snapshot v {{{ γ p, RET p; is_snapshot N γ p ∗ snapshot_content γ v }}}; + read_spec N γ p : + is_snapshot N γ p -∗ + <<< ∀ v : val, snapshot_content γ v >>> + read p @ ⊤∖↑N - <<< pair_content γ (v, v2), RET #() >>>; - writeY_spec N γ (v: val) p : - is_pair N γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - writeY (p, v)%V + <<< snapshot_content γ v, RET v >>>; + write_spec N γ (v: val) p : + is_snapshot N γ p -∗ + <<< ∀ w : val, snapshot_content γ w >>> + write p v @ ⊤∖↑N - <<< pair_content γ (v1, v), RET #() >>>; - readPair_spec N γ p : - is_pair N γ p -∗ - <<< ∀ v1 v2 : val, pair_content γ (v1, v2) >>> - readPair p + <<< snapshot_content γ v, RET #() >>>; + read_with_spec N γ p (l : loc) : + is_snapshot N γ p -∗ + <<< ∀ v w : val, snapshot_content γ v ∗ l ↦ w >>> + read_with p #l @ ⊤∖↑N - <<< pair_content γ (v1, v2), RET (v1, v2) >>>; + <<< snapshot_content γ v ∗ l ↦ w, RET (v, w) >>>; }. Arguments atomic_snapshot _ {_}. Existing Instances - is_pair_persistent pair_content_timeless + is_snapshot_persistent snapshot_content_timeless name_countable name_eqdec. diff --git a/theories/logatom/treiber.v b/theories/logatom/treiber.v index 3eae8c7d34122784d2c83042eaf476f14e26b88e..f69762ae2e577e06ccc116cd1f5415b8a51b5b4f 100644 --- a/theories/logatom/treiber.v +++ b/theories/logatom/treiber.v @@ -108,18 +108,18 @@ Section proof. iDestruct "Hxs" as (hd) "[Hs Hhd]". wp_load. iMod ("Hvs'" with "[Hs Hhd]") as "HP"; first by eauto with iFrame. iModIntro. wp_let. wp_alloc l as "Hl". wp_let. - wp_bind (CAS _ _ _)%E. + wp_bind (CmpXchg _ _ _)%E. iMod "HP" as (xs') "[Hxs' Hvs']". iDestruct "Hxs'" as (hd') "[Hs' Hhd']". destruct (decide (hd = hd')) as [->|Hneq]. - * wp_cas_suc. iDestruct "Hvs'" as "[_ Hvs']". + * wp_cmpxchg_suc. iDestruct "Hvs'" as "[_ Hvs']". iMod ("Hvs'" with "[-]") as "HQ". { simpl. by eauto 10 with iFrame. } - iModIntro. wp_if. eauto. - * wp_cas_fail. + iModIntro. wp_pures. eauto. + * wp_cmpxchg_fail. iDestruct "Hvs'" as "[Hvs' _]". iMod ("Hvs'" with "[-]") as "HP"; first by eauto with iFrame. - iModIntro. wp_if. by iApply "IH". + iModIntro. wp_pures. by iApply "IH". Qed. Lemma pop_atomic_spec (s: loc) : @@ -147,11 +147,11 @@ Section proof. iMod ("Hvs'" with "[-Hhd1 Hhd2 Hxs1]") as "HP". { eauto with iFrame. } iModIntro. wp_let. wp_load. wp_match. wp_proj. - wp_bind (CAS _ _ _). + wp_bind (CmpXchg _ _ _). iMod "HP" as (xs'') "[Hxs'' Hvs']". iDestruct "Hxs''" as (hd'') "[Hs'' Hhd'']". destruct (decide (hd = hd'')) as [->|Hneq]. - + wp_cas_suc. iDestruct "Hvs'" as "[_ Hvs']". + + wp_cmpxchg_suc. iDestruct "Hvs'" as "[_ Hvs']". destruct xs'' as [|x'' xs'']. { simpl. iDestruct "Hhd''" as (?) "H". iExFalso. by iDestruct (@mapsto_agree with "[$Hhd1] [$H]") as %?. } @@ -159,9 +159,9 @@ Section proof. iDestruct (@mapsto_agree with "[$Hhd1] [$Hhd'']") as %[=]. subst. iMod ("Hvs'" with "[-]") as "HQ". { eauto with iFrame. } - iModIntro. wp_if. wp_pures. eauto. - + wp_cas_fail. iDestruct "Hvs'" as "[Hvs' _]". + iModIntro. wp_pures. eauto. + + wp_cmpxchg_fail. iDestruct "Hvs'" as "[Hvs' _]". iMod ("Hvs'" with "[-]") as "HP"; first by eauto with iFrame. - iModIntro. wp_if. by iApply "IH". + iModIntro. wp_pures. by iApply "IH". Qed. End proof. diff --git a/theories/logatom/treiber2.v b/theories/logatom/treiber2.v index 2501052d0396352e570cd87fb8ccdc97a88f1d3f..50af9322ac15e65c52006cf2db67887f10abf091 100644 --- a/theories/logatom/treiber2.v +++ b/theories/logatom/treiber2.v @@ -58,10 +58,10 @@ Definition pop_stack : val := (** * Definition of the required camera *************************************) Class stackG Σ := StackG { - stack_tokG :> inG Σ (authR (optionUR (exclR (listC valC)))) }. + stack_tokG :> inG Σ (authR (optionUR (exclR (listO valO)))) }. Definition stackΣ : gFunctors := - #[GFunctor (authR (optionUR (exclR (listC valC))))]. + #[GFunctor (authR (optionUR (exclR (listO valO))))]. Instance subG_stackΣ {Σ} : subG stackΣ Σ → stackG Σ. Proof. solve_inG. Qed. @@ -260,13 +260,13 @@ Proof. (* We can now perform the load, and get rid of the modality. *) wp_load. iModIntro. iSplitL "Hl Hγ● HPhys"; first by eauto 10 with iFrame. (* We then resume stepping through the program, and focus on the CAS. *) - wp_alloc r as "Hr". wp_inj. wp_bind (CAS _ _ _)%E. + wp_alloc r as "Hr". wp_inj. wp_bind (CmpXchg _ _ _)%E. (* Now, we need to use the invariant again to gain information on [ℓ]. *) iInv "HInv" as (ys) ">[HPhys Hγ●]". iDestruct "HPhys" as (u) "[Hl HPhys]". (* We now reason by case on the success/failure of the CAS. *) destruct (decide (u = w)) as [[= ->]|NE]. - (* The CAS succeeded. *) - wp_cas_suc. { case w; done. (* Administrative stuff. *) } + wp_cmpxchg_suc. { case w; left; done. (* Administrative stuff. *) } (* This was the linearization point. We access the preconditon. *) iMod "AU" as (zs) "[Hγ◯ [_ HClose]]". (* Use agreement on ressource [γ] to learn [zs = ys]. *) @@ -278,14 +278,14 @@ Proof. (* We can eliminate the modality. *) iModIntro. iSplitR "H"; first by eauto 10 with iFrame. (* And conclude the proof easily, after some computation steps. *) - wp_if. iExact "H". + wp_pures. iExact "H". - (* The CAS failed. *) - wp_cas_fail. { eapply not_inj. done. } + wp_cmpxchg_fail. { exact: not_inj. } { case u, w; simpl; eauto. (* Administrative stuff. *) } (* We can eliminate the modality. *) iModIntro. iSplitL "Hγ● Hl HPhys"; first by eauto 10 with iFrame. (* And conclude the proof by induction hypothesis. *) - wp_if. iApply "IH". iExact "AU". + wp_pures. iApply "IH". iExact "AU". Qed. (** As for [push_stack] the specification of [pop_stack] depends on a location @@ -315,7 +315,7 @@ Proof. iModIntro. iSplitL "Hl Hγ● HPhys1"; first by eauto 10 with iFrame. (* We continue stepping through the program, and focus on the CAS. *) wp_let. wp_match. iDestruct "HPhys2" as (r q) "[Hw HP]". - wp_load. wp_let. wp_proj. wp_bind (CAS _ _ _)%E. + wp_load. wp_let. wp_proj. wp_bind (CmpXchg _ _ _)%E. (* We need to use the invariant again to gain information on [ℓ]. *) iInv "HInv" as (ys) ">[H Hγ●]". unfold phys_stack. iDestruct "H" as (u) "[Hl HPhys]". @@ -324,7 +324,7 @@ Proof. (* We reason by case on the success/failure of the CAS. *) destruct (decide (u = Some w)) as [[= ->]|Hx]. * (* The CAS succeeded, so this is the linearization point. *) - wp_cas_suc; first done. + wp_cmpxchg_suc. (* The list [ys] must be non-empty, otherwise the proof is trivial. *) destruct ys; first done. (* We access the precondition, prior to performing an update. *) @@ -341,13 +341,13 @@ Proof. (* Eliminate the modality. *) iModIntro. iSplitR "HΦ"; first by eauto 10 with iFrame. (* And conclude the proof. *) - wp_if. wp_proj. wp_inj. iExact "HΦ". + wp_pures. iExact "HΦ". * (* The CAS failed. *) - wp_cas_fail. { intro H. apply Hx. destruct u; inversion H; done. } + wp_cmpxchg_fail. { intro H. apply Hx. destruct u; inversion H; done. } (* We can eliminate the modality. *) iModIntro. iSplitR "AU"; first by eauto 10 with iFrame. (* And conclude the proof using the induction hypothesis. *) - wp_if. iApply "IH". iExact "AU". + wp_pures. iApply "IH". iExact "AU". - (* The stack is empty, the load was the linearization point, we can hence commit (without updating the stack). So we access the precondition. *) iClear "HPhys1 HPhys2". iMod "AU" as (xs) "[Hγ◯ [_ HClose]]". @@ -358,7 +358,7 @@ Proof. (* We finally eliminate the modality and conclude the proof by taking some computation steps, and using our hypothesis. *) iModIntro. iSplitR "HΦ"; first by eauto 10 with iFrame. - wp_let. wp_match. wp_inj. iExact "HΦ". + wp_pures. iExact "HΦ". Qed. End treiber_stack. diff --git a/theories/logrel/F_mu/fundamental.v b/theories/logrel/F_mu/fundamental.v index 52de92150d1e83edc7f4a02af9eb6392e8495d5a..c6a0c734218d582e055407628e715b0143266c4f 100644 --- a/theories/logrel/F_mu/fundamental.v +++ b/theories/logrel/F_mu/fundamental.v @@ -11,7 +11,7 @@ Notation "Γ ⊨ e : τ" := (log_typed Γ e τ) (at level 74, e, τ at next leve Section fundamental. Context `{irisG F_mu_lang Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Local Tactic Notation "smart_wp_bind" uconstr(ctx) ident(v) constr(Hv) uconstr(Hp) := iApply (wp_bind (fill[ctx])); diff --git a/theories/logrel/F_mu/lang.v b/theories/logrel/F_mu/lang.v index 3f901aa493decf671db51d66530cbffcfc71e7fd..a484ed711ac9e0b143ed7a3a63d2247064fb7d72 100644 --- a/theories/logrel/F_mu/lang.v +++ b/theories/logrel/F_mu/lang.v @@ -167,9 +167,9 @@ Module F_mu. fill_item_val, fill_item_no_val_inj, head_ctx_step_val. Qed. - Canonical Structure stateC := leibnizC state. - Canonical Structure valC := leibnizC val. - Canonical Structure exprC := leibnizC expr. + Canonical Structure stateO := leibnizO state. + Canonical Structure valO := leibnizO val. + Canonical Structure exprO := leibnizO expr. End F_mu. (** Language *) diff --git a/theories/logrel/F_mu/logrel.v b/theories/logrel/F_mu/logrel.v index ff721cbbc3688bc2bfe676b557e8440972f992cf..e5d014c6faf68d37215c46cd872c84c5a1045d87 100644 --- a/theories/logrel/F_mu/logrel.v +++ b/theories/logrel/F_mu/logrel.v @@ -7,57 +7,57 @@ Import uPred. (** interp : is a unary logical relation. *) Section logrel. Context `{irisG F_mu_lang Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Implicit Types τi : D. - Implicit Types Δ : listC D. - Implicit Types interp : listC D → D. + Implicit Types Δ : listO D. + Implicit Types interp : listO D → D. - Program Definition ctx_lookup (x : var) : listC D -n> D := λne Δ, + Program Definition ctx_lookup (x : var) : listO D -n> D := λne Δ, from_option id (cconst True)%I (Δ !! x). Solve Obligations with solve_proper. - Definition interp_unit : listC D -n> D := λne Δ w, (⌜w = UnitV⌝)%I. + Definition interp_unit : listO D -n> D := λne Δ w, (⌜w = UnitV⌝)%I. Program Definition interp_prod - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (∃ w1 w2, ⌜w = PairV w1 w2⌝ ∧ interp1 Δ w1 ∧ interp2 Δ w2)%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_sum - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, ((∃ w1, ⌜w = InjLV w1⌝ ∧ interp1 Δ w1) ∨ (∃ w2, ⌜w = InjRV w2⌝ ∧ interp2 Δ w2))%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_arrow - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ v, interp1 Δ v → WP App (# w) (# v) {{ interp2 Δ }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_forall - (interp : listC D -n> D) : listC D -n> D := λne Δ w, + (interp : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ τi : D, ⌜∀ v, Persistent (τi v)⌝ → WP TApp (# w) {{ interp (τi :: Δ) }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Definition interp_rec1 - (interp : listC D -n> D) (Δ : listC D) (τi : D) : D := λne w, + (interp : listO D -n> D) (Δ : listO D) (τi : D) : D := λne w, (□ (∃ v, ⌜w = FoldV v⌝ ∧ ▷ interp (τi :: Δ) v))%I. Global Instance interp_rec1_contractive - (interp : listC D -n> D) (Δ : listC D) : Contractive (interp_rec1 interp Δ). + (interp : listO D -n> D) (Δ : listO D) : Contractive (interp_rec1 interp Δ). Proof. by solve_contractive. Qed. - Lemma fixpoint_interp_rec1_eq (interp : listC D -n> D) Δ x : + Lemma fixpoint_interp_rec1_eq (interp : listO D -n> D) Δ x : fixpoint (interp_rec1 interp Δ) x ≡ interp_rec1 interp Δ (fixpoint (interp_rec1 interp Δ)) x. Proof. exact: (fixpoint_unfold (interp_rec1 interp Δ) x). Qed. - Program Definition interp_rec (interp : listC D -n> D) : listC D -n> D := λne Δ, + Program Definition interp_rec (interp : listO D -n> D) : listO D -n> D := λne Δ, fixpoint (interp_rec1 interp Δ). Next Obligation. intros interp n Δ1 Δ2 HΔ; apply fixpoint_ne => τi w. solve_proper. Qed. - Fixpoint interp (τ : type) : listC D -n> D := + Fixpoint interp (τ : type) : listO D -n> D := match τ return _ with | TUnit => interp_unit | TProd τ1 τ2 => interp_prod (interp τ1) (interp τ2) @@ -70,11 +70,11 @@ Section logrel. Notation "⟦ τ ⟧" := (interp τ). Definition interp_env (Γ : list type) - (Δ : listC D) (vs : list val) : iProp Σ := + (Δ : listO D) (vs : list val) : iProp Σ := (⌜length Γ = length vs⌝ ∗ [∗] zip_with (λ τ, ⟦ τ ⟧ Δ) Γ vs)%I. Notation "⟦ Γ ⟧*" := (interp_env Γ). - Definition interp_expr (τ : type) (Δ : listC D) (e : expr) : iProp Σ := + Definition interp_expr (τ : type) (Δ : listO D) (e : expr) : iProp Σ := WP e {{ ⟦ τ ⟧ Δ }}%I. Class env_Persistent Δ := @@ -116,7 +116,7 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). Qed. @@ -135,11 +135,11 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. case EQ: (x - length Δ1) => [|n]; simpl. { symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). } - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). Qed. diff --git a/theories/logrel/F_mu/soundness.v b/theories/logrel/F_mu/soundness.v index ade3e7d015244840d2b11f45d4ca04c7ddc07473..ce7b3fbe6ce9bf0a8663770c12c765ad215f8136 100644 --- a/theories/logrel/F_mu/soundness.v +++ b/theories/logrel/F_mu/soundness.v @@ -9,7 +9,7 @@ Theorem soundness Σ `{invPreG Σ} e τ e' thp σ σ' : Proof. intros Hlog ??. cut (adequate NotStuck e σ (λ _ _, True)); first (intros [_ ?]; eauto). eapply (wp_adequacy Σ); eauto. - iIntros (Hinv ?). iModIntro. iExists (λ _ _, True%I). iSplit=> //. + iIntros (Hinv ?). iModIntro. iExists (λ _ _, True%I), (λ _, True%I). iSplit=> //. replace e with e.[env_subst[]] by by asimpl. set (HΣ := IrisG _ _ Hinv (λ _ _ _, True)%I (λ _, True)%I). iApply (wp_wand with "[]"). iApply Hlog; eauto. by iApply interp_env_nil. auto. diff --git a/theories/logrel/F_mu_ref/fundamental.v b/theories/logrel/F_mu_ref/fundamental.v index 8fa60ef1650b67a130afa4167b3aa079a657235f..08187814b9aa62bfdffb45589d9b5cac9fe40435 100644 --- a/theories/logrel/F_mu_ref/fundamental.v +++ b/theories/logrel/F_mu_ref/fundamental.v @@ -10,7 +10,7 @@ Notation "Γ ⊨ e : τ" := (log_typed Γ e τ) (at level 74, e, τ at next leve Section fundamental. Context `{heapG Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Local Tactic Notation "smart_wp_bind" uconstr(ctx) ident(v) constr(Hv) uconstr(Hp) := iApply (wp_bind (fill [ctx])); diff --git a/theories/logrel/F_mu_ref/fundamental_binary.v b/theories/logrel/F_mu_ref/fundamental_binary.v index defc2f0813826ee9c7471ac345c54dcf76cafe9b..e4f49b9884764ce4f740b522c939c535f5c4f3d0 100644 --- a/theories/logrel/F_mu_ref/fundamental_binary.v +++ b/theories/logrel/F_mu_ref/fundamental_binary.v @@ -6,7 +6,7 @@ From iris_examples.logrel.F_mu_ref Require Import rules_binary. Section bin_log_def. Context `{heapG Σ,cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Definition bin_log_related (Γ : list type) (e e' : expr) (τ : type) := ∀ Δ vvs (ρ : cfg F_mu_ref_lang), env_Persistent Δ → @@ -19,9 +19,9 @@ Notation "Γ ⊨ e '≤log≤' e' : τ" := Section fundamental. Context `{heapG Σ,cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Implicit Types e : expr. - Implicit Types Δ : listC D. + Implicit Types Δ : listO D. Hint Resolve to_of_val. Local Tactic Notation "smart_wp_bind" uconstr(ctx) ident(v) ident(w) diff --git a/theories/logrel/F_mu_ref/lang.v b/theories/logrel/F_mu_ref/lang.v index 596874275e030a1d0dd22909d83a93d079dd89b0..49e05629232d6782799036ac5026078311cce50f 100644 --- a/theories/logrel/F_mu_ref/lang.v +++ b/theories/logrel/F_mu_ref/lang.v @@ -214,9 +214,9 @@ Module F_mu_ref. fill_item_val, fill_item_no_val_inj, head_ctx_step_val. Qed. - Canonical Structure stateC := leibnizC state. - Canonical Structure valC := leibnizC val. - Canonical Structure exprC := leibnizC expr. + Canonical Structure stateO := leibnizO state. + Canonical Structure valO := leibnizO val. + Canonical Structure exprO := leibnizO expr. End F_mu_ref. Canonical Structure F_mu_ref_ectxi_lang := EctxiLanguage F_mu_ref.lang_mixin. diff --git a/theories/logrel/F_mu_ref/logrel.v b/theories/logrel/F_mu_ref/logrel.v index 35dc3f0f08e3df3ea1d0fb22e17d03cb3089a8c1..60fd607995d40efbc82724d4c422afb606bc0cca 100644 --- a/theories/logrel/F_mu_ref/logrel.v +++ b/theories/logrel/F_mu_ref/logrel.v @@ -9,52 +9,52 @@ Definition logN : namespace := nroot .@ "logN". (** interp : is a unary logical relation. *) Section logrel. Context `{heapG Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Implicit Types τi : D. - Implicit Types Δ : listC D. - Implicit Types interp : listC D → D. + Implicit Types Δ : listO D. + Implicit Types interp : listO D → D. - Program Definition ctx_lookup (x : var) : listC D -n> D := λne Δ, + Program Definition ctx_lookup (x : var) : listO D -n> D := λne Δ, from_option id (cconst True)%I (Δ !! x). Solve Obligations with solve_proper. - Definition interp_unit : listC D -n> D := λne Δ w, ⌜w = UnitV⌝%I. + Definition interp_unit : listO D -n> D := λne Δ w, ⌜w = UnitV⌝%I. Program Definition interp_prod - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (∃ w1 w2, ⌜w = PairV w1 w2⌝ ∧ interp1 Δ w1 ∧ interp2 Δ w2)%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_sum - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, ((∃ w1, ⌜w = InjLV w1⌝ ∧ interp1 Δ w1) ∨ (∃ w2, ⌜w = InjRV w2⌝ ∧ interp2 Δ w2))%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_arrow - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ v, interp1 Δ v → WP App (# w) (# v) {{ interp2 Δ }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_forall - (interp : listC D -n> D) : listC D -n> D := λne Δ w, + (interp : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ τi : D, ⌜(∀ v, Persistent (τi v))⌝ → WP TApp (# w) {{ interp (τi :: Δ) }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Definition interp_rec1 - (interp : listC D -n> D) (Δ : listC D) (τi : D) : D := λne w, + (interp : listO D -n> D) (Δ : listO D) (τi : D) : D := λne w, (□ (∃ v, ⌜w = FoldV v⌝ ∧ ▷ interp (τi :: Δ) v))%I. Global Instance interp_rec1_contractive - (interp : listC D -n> D) (Δ : listC D) : Contractive (interp_rec1 interp Δ). + (interp : listO D -n> D) (Δ : listO D) : Contractive (interp_rec1 interp Δ). Proof. by solve_contractive. Qed. - Lemma fixpoint_interp_rec1_eq (interp : listC D -n> D) Δ x : + Lemma fixpoint_interp_rec1_eq (interp : listO D -n> D) Δ x : fixpoint (interp_rec1 interp Δ) x ≡ interp_rec1 interp Δ (fixpoint (interp_rec1 interp Δ)) x. Proof. exact: (fixpoint_unfold (interp_rec1 interp Δ) x). Qed. - Program Definition interp_rec (interp : listC D -n> D) : listC D -n> D := λne Δ, + Program Definition interp_rec (interp : listO D -n> D) : listO D -n> D := λne Δ, fixpoint (interp_rec1 interp Δ). Next Obligation. intros interp n Δ1 Δ2 HΔ; apply fixpoint_ne => τi w. solve_proper. @@ -65,11 +65,11 @@ Section logrel. Solve Obligations with solve_proper. Program Definition interp_ref - (interp : listC D -n> D) : listC D -n> D := λne Δ w, + (interp : listO D -n> D) : listO D -n> D := λne Δ w, (∃ l, ⌜w = LocV l⌝ ∧ inv (logN .@ l) (interp_ref_inv l (interp Δ)))%I. Solve Obligations with solve_proper. - Fixpoint interp (τ : type) : listC D -n> D := + Fixpoint interp (τ : type) : listO D -n> D := match τ return _ with | TUnit => interp_unit | TProd τ1 τ2 => interp_prod (interp τ1) (interp τ2) @@ -83,11 +83,11 @@ Section logrel. Notation "⟦ τ ⟧" := (interp τ). Definition interp_env (Γ : list type) - (Δ : listC D) (vs : list val) : iProp Σ := + (Δ : listO D) (vs : list val) : iProp Σ := (⌜length Γ = length vs⌝ ∗ [∗] zip_with (λ τ, ⟦ τ ⟧ Δ) Γ vs)%I. Notation "⟦ Γ ⟧*" := (interp_env Γ). - Definition interp_expr (τ : type) (Δ : listC D) (e : expr) : iProp Σ := + Definition interp_expr (τ : type) (Δ : listO D) (e : expr) : iProp Σ := WP e {{ ⟦ τ ⟧ Δ }}%I. Class env_Persistent Δ := @@ -129,7 +129,7 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). - intros w; simpl; properness; auto. by apply IHτ. @@ -149,11 +149,11 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. case EQ: (x - length Δ1) => [|n]; simpl. { symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). } - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). - intros w; simpl; properness; auto. by apply IHτ. diff --git a/theories/logrel/F_mu_ref/logrel_binary.v b/theories/logrel/F_mu_ref/logrel_binary.v index 48908fe5245b3f61a55eed82325b27ee87f6e67f..1f24a7efa5b280e8f1a3fd76174d27f36eb1447a 100644 --- a/theories/logrel/F_mu_ref/logrel_binary.v +++ b/theories/logrel/F_mu_ref/logrel_binary.v @@ -25,12 +25,12 @@ Definition logN : namespace := nroot .@ "logN". (** interp : is a unary logical relation. *) Section logrel. Context `{heapG Σ, cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Implicit Types τi : D. - Implicit Types Δ : listC D. - Implicit Types interp : listC D → D. + Implicit Types Δ : listO D. + Implicit Types interp : listO D → D. - Definition interp_expr (τi : listC D -n> D) (Δ : listC D) + Definition interp_expr (τi : listO D -n> D) (Δ : listO D) (ee : expr * expr) : iProp Σ := (∀ K, ⤇ fill K (ee.2) → WP ee.1 {{ v, ∃ v', ⤇ fill K (of_val v') ∗ τi Δ (v, v') }})%I. @@ -38,28 +38,28 @@ Section logrel. Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr. Proof. solve_proper. Qed. - Program Definition ctx_lookup (x : var) : listC D -n> D := λne Δ, + Program Definition ctx_lookup (x : var) : listO D -n> D := λne Δ, from_option id (cconst True)%I (Δ !! x). Solve Obligations with solve_proper. - Program Definition interp_unit : listC D -n> D := λne Δ ww, + Program Definition interp_unit : listO D -n> D := λne Δ ww, (⌜ww.1 = UnitV⌝ ∧ ⌜ww.2 = UnitV⌝)%I. Solve Obligations with solve_proper_alt. Program Definition interp_prod - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, (∃ vv1 vv2, ⌜ww = (PairV (vv1.1) (vv2.1), PairV (vv1.2) (vv2.2))⌝ ∧ interp1 Δ vv1 ∧ interp2 Δ vv2)%I. Solve Obligations with repeat intros ?; simpl; auto_equiv. Program Definition interp_sum - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, ((∃ vv, ⌜ww = (InjLV (vv.1), InjLV (vv.2))⌝ ∧ interp1 Δ vv) ∨ (∃ vv, ⌜ww = (InjRV (vv.1), InjRV (vv.2))⌝ ∧ interp2 Δ vv))%I. Solve Obligations with repeat intros ?; simpl; auto_equiv. Program Definition interp_arrow - (interp1 interp2 : listC D -n> D) : listC D -n> D := + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, (□ ∀ vv, interp1 Δ vv → interp_expr @@ -68,7 +68,7 @@ Section logrel. Solve Obligations with repeat intros ?; simpl; auto_equiv. Program Definition interp_forall - (interp : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp : listO D -n> D) : listO D -n> D := λne Δ ww, (□ ∀ τi, ⌜∀ ww, Persistent (τi ww)⌝ → interp_expr @@ -76,19 +76,19 @@ Section logrel. Solve Obligations with repeat intros ?; simpl; auto_equiv. Program Definition interp_rec1 - (interp : listC D -n> D) (Δ : listC D) (τi : D) : D := λne ww, + (interp : listO D -n> D) (Δ : listO D) (τi : D) : D := λne ww, (□ ∃ vv, ⌜ww = (FoldV (vv.1), FoldV (vv.2))⌝ ∧ ▷ interp (τi :: Δ) vv)%I. Solve Obligations with repeat intros ?; simpl; auto_equiv. Global Instance interp_rec1_contractive - (interp : listC D -n> D) (Δ : listC D) : Contractive (interp_rec1 interp Δ). + (interp : listO D -n> D) (Δ : listO D) : Contractive (interp_rec1 interp Δ). Proof. by solve_contractive. Qed. - Lemma fixpoint_interp_rec1_eq (interp : listC D -n> D) Δ x : + Lemma fixpoint_interp_rec1_eq (interp : listO D -n> D) Δ x : fixpoint (interp_rec1 interp Δ) x ≡ interp_rec1 interp Δ (fixpoint (interp_rec1 interp Δ)) x. Proof. exact: (fixpoint_unfold (interp_rec1 interp Δ) x). Qed. - Program Definition interp_rec (interp : listC D -n> D) : listC D -n> D := λne Δ, + Program Definition interp_rec (interp : listO D -n> D) : listO D -n> D := λne Δ, fixpoint (interp_rec1 interp Δ). Next Obligation. intros interp n Δ1 Δ2 HΔ; apply fixpoint_ne => τi ww. solve_proper. @@ -99,12 +99,12 @@ Section logrel. Solve Obligations with repeat intros ?; simpl; auto_equiv. Program Definition interp_ref - (interp : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp : listO D -n> D) : listO D -n> D := λne Δ ww, (∃ ll, ⌜ww = (LocV (ll.1), LocV (ll.2))⌝ ∧ inv (logN .@ ll) (interp_ref_inv ll (interp Δ)))%I. Solve Obligations with repeat intros ?; simpl; auto_equiv. - Fixpoint interp (τ : type) : listC D -n> D := + Fixpoint interp (τ : type) : listO D -n> D := match τ return _ with | TUnit => interp_unit | TProd τ1 τ2 => interp_prod (interp τ1) (interp τ2) @@ -118,7 +118,7 @@ Section logrel. Notation "⟦ τ ⟧" := (interp τ). Definition interp_env (Γ : list type) - (Δ : listC D) (vvs : list (val * val)) : iProp Σ := + (Δ : listO D) (vvs : list (val * val)) : iProp Σ := (⌜length Γ = length vvs⌝ ∗ [∗] zip_with (λ τ, ⟦ τ ⟧ Δ) Γ vvs)%I. Notation "⟦ Γ ⟧*" := (interp_env Γ). @@ -162,7 +162,7 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - unfold interp_expr. intros ww; simpl; properness; auto. by apply (IHτ (_ :: _)). @@ -183,11 +183,11 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. case EQ: (x - length Δ1) => [|n]; simpl. { symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). } - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - unfold interp_expr. intros ww; simpl; properness; auto. apply (IHτ (_ :: _)). diff --git a/theories/logrel/F_mu_ref/rules_binary.v b/theories/logrel/F_mu_ref/rules_binary.v index 57b846f274e361d2a86aeacc40962ea8bcaf14c8..11b7499aa7ee500498758be85d5544d4cc2f0039 100644 --- a/theories/logrel/F_mu_ref/rules_binary.v +++ b/theories/logrel/F_mu_ref/rules_binary.v @@ -7,7 +7,7 @@ Import uPred. Definition specN := nroot .@ "spec". (** The CMRA for the heap of the specification. *) -Definition cfgUR := prodUR (optionUR (exclR exprC)) (gen_heapUR loc val). +Definition cfgUR := prodUR (optionUR (exclR exprO)) (gen_heapUR loc val). (** The CMRA for the thread pool. *) Class cfgSG Σ := CFGSG { cfg_inG :> inG Σ (authR cfgUR); cfg_name : gname }. diff --git a/theories/logrel/F_mu_ref/soundness.v b/theories/logrel/F_mu_ref/soundness.v index f75f2312c47970d3f6ee8c85c281db4e94f2ab37..fff5f5a440b7dc3b447c2c91d21598ba2456d2c4 100644 --- a/theories/logrel/F_mu_ref/soundness.v +++ b/theories/logrel/F_mu_ref/soundness.v @@ -17,7 +17,7 @@ Proof. eapply (wp_adequacy Σ _); eauto. iIntros (Hinv ?). iMod (gen_heap_init σ) as (Hheap) "Hh". - iModIntro. iExists (λ σ _, gen_heap_ctx σ); iFrame. + iModIntro. iExists (λ σ _, gen_heap_ctx σ), (λ _, True%I); iFrame. set (HeapΣ := (HeapG Σ Hinv Hheap)). iApply (wp_wand with "[]"). - replace e with e.[env_subst[]] by by asimpl. diff --git a/theories/logrel/F_mu_ref/soundness_binary.v b/theories/logrel/F_mu_ref/soundness_binary.v index 00bbf0c8f2e08efacdf50e39c937615866a9593c..b16a847278e2c616c034136c27d55e8c3bf3bd2b 100644 --- a/theories/logrel/F_mu_ref/soundness_binary.v +++ b/theories/logrel/F_mu_ref/soundness_binary.v @@ -25,7 +25,7 @@ Proof. rewrite /to_gen_heap fin_maps.map_fmap_empty. iFrame. } set (HeapΣ := HeapG Σ Hinv Hheap). - iExists (λ σ _, gen_heap_ctx σ); iFrame. + iExists (λ σ _, gen_heap_ctx σ), (λ _, True%I); iFrame. iApply wp_fupd. iApply (wp_wand with "[-]"). - iPoseProof (Hlog _ _ with "[$Hcfg]") as "Hrel". { iApply (@logrel_binary.interp_env_nil Σ HeapΣ). } diff --git a/theories/logrel/F_mu_ref_conc/examples/counter.v b/theories/logrel/F_mu_ref_conc/examples/counter.v index 7c4eaa7562228113f3a3b1f12f065bc0004c9418..2f43b80ff66d286f0d1bef8cc00b34411a8b28c3 100644 --- a/theories/logrel/F_mu_ref_conc/examples/counter.v +++ b/theories/logrel/F_mu_ref_conc/examples/counter.v @@ -35,8 +35,8 @@ Definition FG_counter : expr := Section CG_Counter. Context `{heapIG Σ, cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). - Implicit Types Δ : listC D. + Notation D := (prodO valO valO -n> iProp Σ). + Implicit Types Δ : listO D. (* Coarse-grained increment *) Lemma CG_increment_type x Γ : diff --git a/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v b/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v index 998a8e729e2233d3a3a4ee7ea8ff555c24e2fb6b..93fbed2850492b4ee4b80e31cc53f5b8d5c86225 100644 --- a/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v +++ b/theories/logrel/F_mu_ref_conc/examples/stack/refinement.v @@ -10,8 +10,8 @@ Definition stackN : namespace := nroot .@ "stack". Section Stack_refinement. Context `{heapIG Σ, cfgSG Σ, inG Σ (authR stackUR)}. - Notation D := (prodC valC valC -n> iProp Σ). - Implicit Types Δ : listC D. + Notation D := (prodO valO valO -n> iProp Σ). + Implicit Types Δ : listO D. Lemma FG_CG_counter_refinement : [] ⊨ FG_stack ≤log≤ CG_stack : TForall (TProd (TProd diff --git a/theories/logrel/F_mu_ref_conc/examples/stack/stack_rules.v b/theories/logrel/F_mu_ref_conc/examples/stack/stack_rules.v index 99b1bb001ff0b0273c6d8a17060ae567dc85a905..7607df63de322c6c53900bbc5ed80b73a3ea62a5 100644 --- a/theories/logrel/F_mu_ref_conc/examples/stack/stack_rules.v +++ b/theories/logrel/F_mu_ref_conc/examples/stack/stack_rules.v @@ -5,7 +5,7 @@ Import uPred. From iris.algebra Require deprecated. Import deprecated.dec_agree. -Definition stackUR : ucmraT := gmapUR loc (agreeR valC). +Definition stackUR : ucmraT := gmapUR loc (agreeR valO). Class stackG Σ := StackG { stack_inG :> inG Σ (authR stackUR); stack_name : gname }. @@ -17,7 +17,7 @@ Notation "l ↦ˢᵗᵏ v" := (stack_mapsto l v) (at level 20) : bi_scope. Section Rules. Context `{stackG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Global Instance stack_mapsto_persistent l v : Persistent (l ↦ˢᵗᵏ v). Proof. apply _. Qed. diff --git a/theories/logrel/F_mu_ref_conc/fundamental_binary.v b/theories/logrel/F_mu_ref_conc/fundamental_binary.v index b683aa59d3949624daa9fe8a704b36d61fe752d6..260f457f25904e0ea04e215fa09c507b6b484e38 100644 --- a/theories/logrel/F_mu_ref_conc/fundamental_binary.v +++ b/theories/logrel/F_mu_ref_conc/fundamental_binary.v @@ -6,7 +6,7 @@ From iris.program_logic Require Export lifting. Section bin_log_def. Context `{heapIG Σ, cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Definition bin_log_related (Γ : list type) (e e' : expr) (τ : type) := ∀ Δ vvs, env_Persistent Δ → @@ -19,9 +19,9 @@ Notation "Γ ⊨ e '≤log≤' e' : τ" := Section fundamental. Context `{heapIG Σ, cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Implicit Types e : expr. - Implicit Types Δ : listC D. + Implicit Types Δ : listO D. Hint Resolve to_of_val. Local Tactic Notation "smart_wp_bind" uconstr(ctx) ident(v) ident(w) diff --git a/theories/logrel/F_mu_ref_conc/fundamental_unary.v b/theories/logrel/F_mu_ref_conc/fundamental_unary.v index 97fa81ba433ba487eae6bcec05a259b87593dd6a..e79b8a7351b60793496016dc8f06823ba9709212 100644 --- a/theories/logrel/F_mu_ref_conc/fundamental_unary.v +++ b/theories/logrel/F_mu_ref_conc/fundamental_unary.v @@ -11,7 +11,7 @@ Notation "Γ ⊨ e : τ" := (log_typed Γ e τ) (at level 74, e, τ at next leve Section typed_interp. Context `{heapIG Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Local Tactic Notation "smart_wp_bind" uconstr(ctx) ident(v) constr(Hv) uconstr(Hp) := iApply (wp_bind (fill [ctx])); diff --git a/theories/logrel/F_mu_ref_conc/lang.v b/theories/logrel/F_mu_ref_conc/lang.v index f852b58c5bbfe4cd6f64b50bce19a58145ea26f9..274441f4ec8ce318c920a1ca5954f648ae8c88c6 100644 --- a/theories/logrel/F_mu_ref_conc/lang.v +++ b/theories/logrel/F_mu_ref_conc/lang.v @@ -303,9 +303,9 @@ Module F_mu_ref_conc. fill_item_val, fill_item_no_val_inj, head_ctx_step_val. Qed. - Canonical Structure stateC := leibnizC state. - Canonical Structure valC := leibnizC val. - Canonical Structure exprC := leibnizC expr. + Canonical Structure stateO := leibnizO state. + Canonical Structure valO := leibnizO val. + Canonical Structure exprO := leibnizO expr. End F_mu_ref_conc. Canonical Structure F_mu_ref_conc_ectxi_lang := diff --git a/theories/logrel/F_mu_ref_conc/logrel_binary.v b/theories/logrel/F_mu_ref_conc/logrel_binary.v index 469776613373e6a46e1b106240c4ff2de1c1e374..cdbb64c6be5a4a52698d774a696f1a203ff02e68 100644 --- a/theories/logrel/F_mu_ref_conc/logrel_binary.v +++ b/theories/logrel/F_mu_ref_conc/logrel_binary.v @@ -28,12 +28,12 @@ Definition logN : namespace := nroot .@ "logN". (** interp : is a unary logical relation. *) Section logrel. Context `{heapIG Σ, cfgSG Σ}. - Notation D := (prodC valC valC -n> iProp Σ). + Notation D := (prodO valO valO -n> iProp Σ). Implicit Types τi : D. - Implicit Types Δ : listC D. - Implicit Types interp : listC D → D. + Implicit Types Δ : listO D. + Implicit Types interp : listO D → D. - Definition interp_expr (τi : listC D -n> D) (Δ : listC D) + Definition interp_expr (τi : listO D -n> D) (Δ : listO D) (ee : expr * expr) : iProp Σ := (∀ j K, j ⤇ fill K (ee.2) → WP ee.1 {{ v, ∃ v', j ⤇ fill K (of_val v') ∗ τi Δ (v, v') }})%I. @@ -41,35 +41,35 @@ Section logrel. Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr. Proof. unfold interp_expr; solve_proper. Qed. - Program Definition ctx_lookup (x : var) : listC D -n> D := λne Δ, + Program Definition ctx_lookup (x : var) : listO D -n> D := λne Δ, from_option id (cconst True)%I (Δ !! x). Solve Obligations with solve_proper. - Program Definition interp_unit : listC D -n> D := λne Δ ww, + Program Definition interp_unit : listO D -n> D := λne Δ ww, (⌜ww.1 = UnitV⌝ ∧ ⌜ww.2 = UnitV⌝)%I. Solve Obligations with solve_proper_alt. - Program Definition interp_nat : listC D -n> D := λne Δ ww, + Program Definition interp_nat : listO D -n> D := λne Δ ww, (∃ n : nat, ⌜ww.1 = #nv n⌝ ∧ ⌜ww.2 = #nv n⌝)%I. Solve Obligations with solve_proper. - Program Definition interp_bool : listC D -n> D := λne Δ ww, + Program Definition interp_bool : listO D -n> D := λne Δ ww, (∃ b : bool, ⌜ww.1 = #♭v b⌝ ∧ ⌜ww.2 = #♭v b⌝)%I. Solve Obligations with solve_proper. Program Definition interp_prod - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, (∃ vv1 vv2, ⌜ww = (PairV (vv1.1) (vv2.1), PairV (vv1.2) (vv2.2))⌝ ∧ interp1 Δ vv1 ∧ interp2 Δ vv2)%I. Solve Obligations with solve_proper. Program Definition interp_sum - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, ((∃ vv, ⌜ww = (InjLV (vv.1), InjLV (vv.2))⌝ ∧ interp1 Δ vv) ∨ (∃ vv, ⌜ww = (InjRV (vv.1), InjRV (vv.2))⌝ ∧ interp2 Δ vv))%I. Solve Obligations with solve_proper. Program Definition interp_arrow - (interp1 interp2 : listC D -n> D) : listC D -n> D := + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ ww, (□ ∀ vv, interp1 Δ vv → interp_expr @@ -78,7 +78,7 @@ Section logrel. Solve Obligations with solve_proper. Program Definition interp_forall - (interp : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp : listO D -n> D) : listO D -n> D := λne Δ ww, (□ ∀ τi, ⌜∀ ww, Persistent (τi ww)⌝ → interp_expr @@ -86,19 +86,19 @@ Section logrel. Solve Obligations with solve_proper. Program Definition interp_rec1 - (interp : listC D -n> D) (Δ : listC D) (τi : D) : D := λne ww, + (interp : listO D -n> D) (Δ : listO D) (τi : D) : D := λne ww, (□ ∃ vv, ⌜ww = (FoldV (vv.1), FoldV (vv.2))⌝ ∧ ▷ interp (τi :: Δ) vv)%I. Solve Obligations with solve_proper. Global Instance interp_rec1_contractive - (interp : listC D -n> D) (Δ : listC D) : Contractive (interp_rec1 interp Δ). + (interp : listO D -n> D) (Δ : listO D) : Contractive (interp_rec1 interp Δ). Proof. solve_contractive. Qed. - Lemma fixpoint_interp_rec1_eq (interp : listC D -n> D) Δ x : + Lemma fixpoint_interp_rec1_eq (interp : listO D -n> D) Δ x : fixpoint (interp_rec1 interp Δ) x ≡ interp_rec1 interp Δ (fixpoint (interp_rec1 interp Δ)) x. Proof. exact: (fixpoint_unfold (interp_rec1 interp Δ) x). Qed. - Program Definition interp_rec (interp : listC D -n> D) : listC D -n> D := λne Δ, + Program Definition interp_rec (interp : listO D -n> D) : listO D -n> D := λne Δ, fixpoint (interp_rec1 interp Δ). Next Obligation. intros interp n Δ1 Δ2 HΔ; apply fixpoint_ne => τi ww. solve_proper. @@ -109,12 +109,12 @@ Section logrel. Solve Obligations with solve_proper. Program Definition interp_ref - (interp : listC D -n> D) : listC D -n> D := λne Δ ww, + (interp : listO D -n> D) : listO D -n> D := λne Δ ww, (∃ ll, ⌜ww = (LocV (ll.1), LocV (ll.2))⌝ ∧ inv (logN .@ ll) (interp_ref_inv ll (interp Δ)))%I. Solve Obligations with solve_proper. - Fixpoint interp (τ : type) : listC D -n> D := + Fixpoint interp (τ : type) : listO D -n> D := match τ return _ with | TUnit => interp_unit | TNat => interp_nat @@ -130,7 +130,7 @@ Section logrel. Notation "⟦ τ ⟧" := (interp τ). Definition interp_env (Γ : list type) - (Δ : listC D) (vvs : list (val * val)) : iProp Σ := + (Δ : listO D) (vvs : list (val * val)) : iProp Σ := (⌜length Γ = length vvs⌝ ∗ [∗] zip_with (λ τ, ⟦ τ ⟧ Δ) Γ vvs)%I. Notation "⟦ Γ ⟧*" := (interp_env Γ). @@ -174,7 +174,7 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - unfold interp_expr. intros ww; simpl; properness; auto. by apply (IHτ (_ :: _)). @@ -195,11 +195,11 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. case EQ: (x - length Δ1) => [|n]; simpl. { symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). } - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - unfold interp_expr. intros ww; simpl; properness; auto. apply (IHτ (_ :: _)). diff --git a/theories/logrel/F_mu_ref_conc/logrel_unary.v b/theories/logrel/F_mu_ref_conc/logrel_unary.v index 53715af0032bb367ffa819e9380ef77ebfd436d9..a1da8812dad1a577651b6270dcb802599fc1593b 100644 --- a/theories/logrel/F_mu_ref_conc/logrel_unary.v +++ b/theories/logrel/F_mu_ref_conc/logrel_unary.v @@ -10,53 +10,53 @@ Definition logN : namespace := nroot .@ "logN". (** interp : is a unary logical relation. *) Section logrel. Context `{heapIG Σ}. - Notation D := (valC -n> iProp Σ). + Notation D := (valO -n> iProp Σ). Implicit Types τi : D. - Implicit Types Δ : listC D. - Implicit Types interp : listC D → D. + Implicit Types Δ : listO D. + Implicit Types interp : listO D → D. - Program Definition env_lookup (x : var) : listC D -n> D := λne Δ, + Program Definition env_lookup (x : var) : listO D -n> D := λne Δ, from_option id (cconst True)%I (Δ !! x). Solve Obligations with solve_proper. - Definition interp_unit : listC D -n> D := λne Δ w, ⌜w = UnitV⌝%I. - Definition interp_nat : listC D -n> D := λne Δ w, ⌜∃ n, w = #nv n⌝%I. - Definition interp_bool : listC D -n> D := λne Δ w, ⌜∃ n, w = #♭v n⌝%I. + Definition interp_unit : listO D -n> D := λne Δ w, ⌜w = UnitV⌝%I. + Definition interp_nat : listO D -n> D := λne Δ w, ⌜∃ n, w = #nv n⌝%I. + Definition interp_bool : listO D -n> D := λne Δ w, ⌜∃ n, w = #♭v n⌝%I. Program Definition interp_prod - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (∃ w1 w2, ⌜w = PairV w1 w2⌝ ∧ interp1 Δ w1 ∧ interp2 Δ w2)%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_sum - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, ((∃ w1, ⌜w = InjLV w1⌝ ∧ interp1 Δ w1) ∨ (∃ w2, ⌜w = InjRV w2⌝ ∧ interp2 Δ w2))%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_arrow - (interp1 interp2 : listC D -n> D) : listC D -n> D := λne Δ w, + (interp1 interp2 : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ v, interp1 Δ v → WP App (of_val w) (of_val v) {{ interp2 Δ }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Program Definition interp_forall - (interp : listC D -n> D) : listC D -n> D := λne Δ w, + (interp : listO D -n> D) : listO D -n> D := λne Δ w, (□ ∀ τi : D, ⌜∀ v, Persistent (τi v)⌝ → WP TApp (of_val w) {{ interp (τi :: Δ) }})%I. Solve Obligations with repeat intros ?; simpl; solve_proper. Definition interp_rec1 - (interp : listC D -n> D) (Δ : listC D) (τi : D) : D := λne w, + (interp : listO D -n> D) (Δ : listO D) (τi : D) : D := λne w, (□ (∃ v, ⌜w = FoldV v⌝ ∧ ▷ interp (τi :: Δ) v))%I. Global Instance interp_rec1_contractive - (interp : listC D -n> D) (Δ : listC D) : Contractive (interp_rec1 interp Δ). + (interp : listO D -n> D) (Δ : listO D) : Contractive (interp_rec1 interp Δ). Proof. by solve_contractive. Qed. - Lemma fixpoint_interp_rec1_eq (interp : listC D -n> D) Δ x : + Lemma fixpoint_interp_rec1_eq (interp : listO D -n> D) Δ x : fixpoint (interp_rec1 interp Δ) x ≡ interp_rec1 interp Δ (fixpoint (interp_rec1 interp Δ)) x. Proof. exact: (fixpoint_unfold (interp_rec1 interp Δ) x). Qed. - Program Definition interp_rec (interp : listC D -n> D) : listC D -n> D := λne Δ, + Program Definition interp_rec (interp : listO D -n> D) : listO D -n> D := λne Δ, fixpoint (interp_rec1 interp Δ). Next Obligation. intros interp n Δ1 Δ2 HΔ; apply fixpoint_ne => τi w. solve_proper. @@ -67,11 +67,11 @@ Section logrel. Solve Obligations with solve_proper. Program Definition interp_ref - (interp : listC D -n> D) : listC D -n> D := λne Δ w, + (interp : listO D -n> D) : listO D -n> D := λne Δ w, (∃ l, ⌜w = LocV l⌝ ∧ inv (logN .@ l) (interp_ref_inv l (interp Δ)))%I. Solve Obligations with solve_proper. - Fixpoint interp (τ : type) : listC D -n> D := + Fixpoint interp (τ : type) : listO D -n> D := match τ return _ with | TUnit => interp_unit | TNat => interp_nat @@ -87,11 +87,11 @@ Section logrel. Notation "⟦ τ ⟧" := (interp τ). Definition interp_env (Γ : list type) - (Δ : listC D) (vs : list val) : iProp Σ := + (Δ : listO D) (vs : list val) : iProp Σ := (⌜length Γ = length vs⌝ ∗ [∗] zip_with (λ τ, ⟦ τ ⟧ Δ) Γ vs)%I. Notation "⟦ Γ ⟧*" := (interp_env Γ). - Definition interp_expr (τ : type) (Δ : listC D) (e : expr) : iProp Σ := + Definition interp_expr (τ : type) (Δ : listO D) (e : expr) : iProp Σ := WP e {{ ⟦ τ ⟧ Δ }}%I. Class env_Persistent Δ := @@ -133,7 +133,7 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). - intros w; simpl; properness; auto. by apply IHτ. @@ -152,11 +152,11 @@ Section logrel. - rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl. { by rewrite !lookup_app_l. } (* FIXME: Ideally we wouldn't have to do this kinf of surgery. *) - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. case EQ: (x - length Δ1) => [|n]; simpl. { symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). } - change (bi_ofeC (uPredI (iResUR Σ))) with (uPredC (iResUR Σ)). + change (bi_ofeO (uPredI (iResUR Σ))) with (uPredO (iResUR Σ)). rewrite !lookup_app_r; [|lia ..]. do 2 f_equiv. lia. - intros w; simpl; properness; auto. apply (IHτ (_ :: _)). - intros w; simpl; properness; auto. by apply IHτ. diff --git a/theories/logrel/F_mu_ref_conc/rules_binary.v b/theories/logrel/F_mu_ref_conc/rules_binary.v index 7c2449c084ebff96740364625142ac61b0712831..1d8f47e9aff9b58c1a54ed6e50f175d6f27ceb5f 100644 --- a/theories/logrel/F_mu_ref_conc/rules_binary.v +++ b/theories/logrel/F_mu_ref_conc/rules_binary.v @@ -8,7 +8,7 @@ Import uPred. Definition specN := nroot .@ "spec". (** The CMRA for the heap of the specification. *) -Definition tpoolUR : ucmraT := gmapUR nat (exclR exprC). +Definition tpoolUR : ucmraT := gmapUR nat (exclR exprO). Definition cfgUR := prodUR tpoolUR (gen_heapUR loc val). Fixpoint to_tpool_go (i : nat) (tp : list expr) : tpoolUR := @@ -107,7 +107,7 @@ Section conversions. - by rewrite lookup_insert tpool_lookup lookup_app_r // Nat.sub_diag. - rewrite lookup_insert_ne; last lia. rewrite !tpool_lookup ?lookup_ge_None_2 ?app_length //=; - change (ofe_car exprC) with expr; lia. + change (ofe_car exprO) with expr; lia. Qed. Lemma tpool_singleton_included tp j e : diff --git a/theories/logrel/F_mu_ref_conc/soundness_binary.v b/theories/logrel/F_mu_ref_conc/soundness_binary.v index bf10c67126d424921d277bd7b701e2a9ca4f731a..d747ae174d19922d61ebf5f07bdd46297cdfb8e4 100644 --- a/theories/logrel/F_mu_ref_conc/soundness_binary.v +++ b/theories/logrel/F_mu_ref_conc/soundness_binary.v @@ -22,7 +22,7 @@ Proof. iMod (inv_alloc specN _ (spec_inv ([e'], ∅)) with "[Hcfg1]") as "#Hcfg". { iNext. iExists [e'], ∅. rewrite /to_gen_heap fin_maps.map_fmap_empty. auto. } set (HeapΣ := (HeapIG Σ Hinv Hheap)). - iExists (λ σ _, gen_heap_ctx σ); iFrame. + iExists (λ σ _, gen_heap_ctx σ), (λ _, True%I); iFrame. iApply wp_fupd. iApply wp_wand_r. iSplitL. iPoseProof ((Hlog _ _ [] []) with "[]") as "Hrel". diff --git a/theories/logrel/F_mu_ref_conc/soundness_unary.v b/theories/logrel/F_mu_ref_conc/soundness_unary.v index 7a1213f2bcb096c6e509add2c727acb6cb0e3db0..c122efada7b864d04ce244740b0a79ffe2636135 100644 --- a/theories/logrel/F_mu_ref_conc/soundness_unary.v +++ b/theories/logrel/F_mu_ref_conc/soundness_unary.v @@ -16,7 +16,7 @@ Proof. intros Hlog ??. cut (adequate NotStuck e σ (λ _ _, True)); first (intros [_ ?]; eauto). eapply (wp_adequacy Σ _). iIntros (Hinv ?). iMod (gen_heap_init σ) as (Hheap) "Hh". - iModIntro. iExists (λ σ _, gen_heap_ctx σ); iFrame. + iModIntro. iExists (λ σ _, gen_heap_ctx σ), (λ _, True%I); iFrame. set (HeapΣ := (HeapIG Σ Hinv Hheap)). iApply (wp_wand with "[]"). - replace e with e.[env_subst[]] by by asimpl. diff --git a/theories/logrel/prelude/base.v b/theories/logrel/prelude/base.v index ce6f8bef757c932a8b2b90e8d214edb7c1d64c99..1007bee3497f775f1a6d96bce4a7d4981f53214b 100644 --- a/theories/logrel/prelude/base.v +++ b/theories/logrel/prelude/base.v @@ -5,7 +5,7 @@ From iris.base_logic Require Import invariants. From Autosubst Require Export Autosubst. Import uPred. -Canonical Structure varC := leibnizC var. +Canonical Structure varO := leibnizO var. Section Autosubst_Lemmas. Context {term : Type} {Ids_term : Ids term} diff --git a/theories/logrel/stlc/soundness.v b/theories/logrel/stlc/soundness.v index be359de05dbb77fcaf05f8277ec3d2886fed6fdc..3c50d7e15567850732853a5997bf7be0a59b8933 100644 --- a/theories/logrel/stlc/soundness.v +++ b/theories/logrel/stlc/soundness.v @@ -16,8 +16,7 @@ Proof. set (Σ := invΣ). intros. cut (adequate NotStuck e () (λ _ _, True)); first (intros [_ Hsafe]; eauto). eapply (wp_adequacy Σ _). iIntros (Hinv ?). - iModIntro. iExists (λ _ _, True%I). iSplit=>//. + iModIntro. iExists (λ _ _, True%I), (λ _, True%I). iSplit=>//. set (HΣ := IrisG _ _ Hinv (λ _ _ _, True)%I (λ _, True)%I). iApply (wp_wand with "[]"). by iApply wp_soundness. eauto. Qed. - diff --git a/theories/logrel_heaplang/ltyping.v b/theories/logrel_heaplang/ltyping.v index ed3fb03443609dcac13de08b3a3ce2fb04e412a1..621b62ecdddd7eecad451c4615c409d984eaf6df 100644 --- a/theories/logrel_heaplang/ltyping.v +++ b/theories/logrel_heaplang/ltyping.v @@ -20,11 +20,11 @@ Section lty_ofe. Instance lty_equiv : Equiv (lty Σ) := λ A B, ∀ w, A w ≡ B w. Instance lty_dist : Dist (lty Σ) := λ n A B, ∀ w, A w ≡{n}≡ B w. Lemma lty_ofe_mixin : OfeMixin (lty Σ). - Proof. by apply (iso_ofe_mixin (lty_car : _ → val -c> _)). Qed. + Proof. by apply (iso_ofe_mixin (lty_car : _ → val -d> _)). Qed. Canonical Structure ltyC := OfeT (lty Σ) lty_ofe_mixin. Global Instance lty_cofe : Cofe ltyC. Proof. - apply (iso_cofe_subtype' (λ A : val -c> iProp Σ, ∀ w, Persistent (A w)) + apply (iso_cofe_subtype' (λ A : val -d> iProp Σ, ∀ w, Persistent (A w)) (@Lty _) lty_car)=> //. - apply _. - apply limit_preserving_forall=> w. @@ -160,8 +160,13 @@ Section types_properties. Proof. iIntros (v). by iDestruct 1 as (i ->) "?". Qed. (* Operator typing *) - Global Instance lty_bin_op_eq A : LTyBinOp EqOp A A lty_bool. - Proof. iIntros (v1 v2) "_ _". rewrite /bin_op_eval /lty_car /=. eauto. Qed. + Global Instance lty_bin_op_eq A : LTyUnboxed A → LTyBinOp EqOp A A lty_bool. + Proof. + iIntros (? v1 v2) "A1 _". rewrite /bin_op_eval /lty_car /=. + iDestruct (lty_unboxed with "A1") as %Hunb. + rewrite decide_True; last solve_vals_compare_safe. + eauto. + Qed. Global Instance lty_bin_op_arith op : TCElemOf op [PlusOp; MinusOp; MultOp; QuotOp; RemOp; AndOp; OrOp; XorOp; ShiftLOp; ShiftROp] → @@ -374,18 +379,16 @@ Section types_properties. iInv (tyN.@l) as (v) "[>Hl Hv]"; iDestruct "Hv" as (n') "> ->". wp_faa. iModIntro. eauto 10. Qed. - Lemma ltyped_cas Γ A e1 e2 e3 : + Lemma ltyped_cmpxchg Γ A e1 e2 e3 : LTyUnboxed A → - (Γ ⊨ e1 : ref A) -∗ (Γ ⊨ e2 : A) -∗ (Γ ⊨ e3 : A) -∗ Γ ⊨ CAS e1 e2 e3 : lty_bool. + (Γ ⊨ e1 : ref A) -∗ (Γ ⊨ e2 : A) -∗ (Γ ⊨ e3 : A) -∗ Γ ⊨ CmpXchg e1 e2 e3 : A * lty_bool. Proof. intros. iIntros "#H1 #H2 #H3" (vs) "!# #HΓ /=". wp_apply (wp_wand with "(H3 [//])"); iIntros (w3) "HA3". wp_apply (wp_wand with "(H2 [//])"); iIntros (w2) "HA2". iDestruct (lty_unboxed with "HA2") as %?. wp_apply (wp_wand with "(H1 [//])"); iIntros (w1); iDestruct 1 as (l ->) "#?". - iInv (tyN.@l) as (v) "[>Hl Hv]". - destruct (decide (v = w2)) as [->|]. - - wp_cas_suc. eauto 10. - - wp_cas_fail. eauto 10. + iInv (tyN.@l) as (v) "[>Hl #Hv]". wp_cmpxchg as ?|?; iModIntro; + (iSplitL; [by eauto 12 with iFrame | iExists _, _; eauto]). Qed. End types_properties. diff --git a/theories/spanning_tree/mon.v b/theories/spanning_tree/mon.v index 9e14e57adf867ea9e3c72762ff72ae3b9e3f5601..efe6069f2cd799ff30410fda077d33f7a7339977 100644 --- a/theories/spanning_tree/mon.v +++ b/theories/spanning_tree/mon.v @@ -11,11 +11,11 @@ From stdpp Require Import gmap mapset. From iris_examples.spanning_tree Require Import graph. (* children cofe *) -Canonical Structure chlC := leibnizC (option loc * option loc). +Canonical Structure chlO := leibnizO (option loc * option loc). (* The graph monoid. *) Definition graphN : namespace := nroot .@ "SPT_graph". Definition graphUR : ucmraT := - optionUR (prodR fracR (gmapR loc (exclR chlC))). + optionUR (prodR fracR (gmapR loc (exclR chlO))). (* The monoid for talking about which nodes are marked. These markings are duplicatable. *) Definition markingUR : ucmraT := gsetUR loc. @@ -67,9 +67,9 @@ Section marking_definitions. End marking_definitions. (* The monoid representing graphs *) -Definition Gmon := gmapR loc (exclR chlC). +Definition Gmon := gmapR loc (exclR chlO). -Definition excl_chlC_chl (ch : exclR chlC) : option (option loc * option loc) := +Definition excl_chlC_chl (ch : exclR chlO) : option (option loc * option loc) := match ch with | Excl w => Some w | Excl_Bot => None @@ -123,7 +123,7 @@ Definition of_graph_empty (g : graph loc) : of_graph g ∅ = fmap (λ x, (false, x)) g. Proof. apply: map_eq => i. - rewrite lookup_imap /of_graph_elem lookup_fmap lookup_omap //. + rewrite map_lookup_imap /of_graph_elem lookup_fmap lookup_omap //. Qed. Lemma of_graph_dom_eq g G : @@ -133,7 +133,7 @@ Proof. intros HGvl. rewrite Gmon_graph_dom // => Hd. apply map_eq => i. assert (Hd' : i ∈ dom (gset _) g ↔ i ∈ dom (gset _) G) by (by rewrite Hd). revert Hd'; clear Hd. specialize (HGvl i); revert HGvl. - rewrite /of_graph /of_graph_elem /Gmon_graph lookup_imap lookup_fmap + rewrite /of_graph /of_graph_elem /Gmon_graph map_lookup_imap lookup_fmap lookup_omap ?elem_of_dom. case _ : (g !! i); case _ : (G !! i) => [[]|] /=; inversion 1; eauto; intros [? ?]; @@ -231,7 +231,7 @@ Proof. intros H. by rewrite lookup_op lookup_singleton_ne //= left_id_L. Qed. Lemma of_graph_dom g G : dom (gset loc) (of_graph g G) = dom (gset _) g. Proof. apply elem_of_equiv_L=>i. - rewrite ?elem_of_dom lookup_imap /of_graph_elem lookup_omap. + rewrite ?elem_of_dom map_lookup_imap /of_graph_elem lookup_omap. case _ : (g !! i) => [x|]; case _ : (G !! i) => [[]|] //=; split; intros [? Hcn]; inversion Hcn; eauto. Qed. @@ -239,7 +239,7 @@ Qed. Lemma in_dom_of_graph (g : graph loc) (G : Gmon) x (b : bool) v : ✓ G → of_graph g G !! x = Some (b, v) → b ↔ x ∈ dom (gset _) G. Proof. - rewrite /of_graph /of_graph_elem lookup_imap lookup_omap elem_of_dom. + rewrite /of_graph /of_graph_elem map_lookup_imap lookup_omap elem_of_dom. intros Hvl; specialize (Hvl x); revert Hvl; case _ : (g !! x) => [?|]; case _ : (G !! x) => [[] ?|] //=; intros Hvl; inversion Hvl; try (inversion 1; subst); split; @@ -256,14 +256,14 @@ Lemma mark_update_lookup (g : graph loc) (G : Gmon) x v : ✓ ((x [↦] v) ⋅ G) → of_graph g ((x [↦] v) ⋅ G) !! x = Some (true, v). Proof. rewrite elem_of_dom /is_Some. intros [w H1] H2. - rewrite /of_graph /of_graph_elem lookup_imap H1 lookup_omap; simpl. + rewrite /of_graph /of_graph_elem map_lookup_imap H1 lookup_omap; simpl. rewrite mark_update_lookup_base; trivial. Qed. Lemma mark_update_lookup_ne (g : graph loc) (G : Gmon) x i v : i ≠ x → of_graph g ((x [↦] v) ⋅ G) !! i = (of_graph g G) !! i. Proof. - intros H. rewrite /of_graph /of_graph_elem ?lookup_imap ?lookup_omap; simpl. + intros H. rewrite /of_graph /of_graph_elem ?map_lookup_imap ?lookup_omap; simpl. rewrite mark_update_lookup_ne_base //=. Qed. @@ -419,7 +419,7 @@ Lemma delete_marked g G x w : Proof. apply: map_eq => i. destruct (decide (i = x)). - subst; by rewrite ?lookup_delete. - - rewrite ?lookup_delete_ne //= /of_graph /of_graph_elem ?lookup_imap + - rewrite ?lookup_delete_ne //= /of_graph /of_graph_elem ?map_lookup_imap ?lookup_omap; case _ : (g !! i) => [v|] //=. by rewrite lookup_op lookup_singleton_ne //= left_id_L. Qed. @@ -533,7 +533,7 @@ Lemma graph_op_path' (G G' : Gmon) x z p : ✓ (G ⋅ G') → x ∈ dom (gset _) G' → valid_path (Gmon_graph G) z x p → False. Proof. rewrite comm; apply graph_op_path. Qed. -Lemma in_dom_singleton (x : loc) (w : chlC) : +Lemma in_dom_singleton (x : loc) (w : chlO) : x ∈ dom (gset loc) (x [↦] w : gmap loc _). Proof. by rewrite dom_singleton elem_of_singleton. Qed. @@ -705,7 +705,7 @@ Qed. Lemma of_graph_unmarked (g : graph loc) (G : Gmon) x v : of_graph g G !! x = Some (false, v) → g !! x = Some v. Proof. - rewrite lookup_imap /of_graph_elem lookup_omap. + rewrite map_lookup_imap /of_graph_elem lookup_omap. case _ : (g !! x); case _ : (G !! x) => [[]|]; by inversion 1. Qed. Lemma get_lr_disj (G G' : Gmon) i : ✓ (G ⋅ G') → diff --git a/theories/spanning_tree/spanning.v b/theories/spanning_tree/spanning.v index 6afd1516514cffeca605b39dce8160cd965a9b62..32e75bbffd6687beacea983826fae875ec4bc8d5 100644 --- a/theories/spanning_tree/spanning.v +++ b/theories/spanning_tree/spanning.v @@ -62,6 +62,7 @@ Section Helpers. { iNext. unfold graph_inv at 2. iExists _; iFrame; auto. } iModIntro. wp_proj. wp_let. destruct u as [u1 u2]; simpl. + wp_bind (CmpXchg _ _ _). iMod (cinv_open _ graphN with "Hgr key") as "(>Hinv & key & Hclose)"; first done. unfold graph_inv at 2. @@ -75,7 +76,7 @@ Section Helpers. rewrite Hil2' in Hil2; inversion Hil2; subst. iDestruct (auth_own_graph_valid with "Hi1") as %Hvl. destruct u as [[] uch]. - - wp_cas_fail; first done. + - wp_cmpxchg_fail. iDestruct (graph_close with "[Hi3 Hil3 Hil4]") as "Hi3"; eauto. { iFrame. iExists _; eauto. iSplitR; eauto. by iExists _; iFrame. } @@ -83,9 +84,9 @@ Section Helpers. { by eapply in_dom_of_graph. } iMod ("Hclose" with "[Hi1 Hi2 Hi3]") as "_". { iNext. unfold graph_inv at 2. iExists _; iFrame; auto. } - iModIntro. iFrame. iRight; by iFrame. + iModIntro. iFrame. wp_pures. iRight; by iFrame. - (* CAS succeeds *) - wp_cas_suc; first done. + wp_cmpxchg_suc. iMod (mark_graph _ _ x uch with "[Hi1 Hx]") as "[Hi1 Hx]"; try by iFrame. { apply (proj1 (not_elem_of_dom (D := gset loc) G' x)). intros Hid. eapply in_dom_of_graph in Hid; eauto; tauto. } @@ -100,7 +101,7 @@ Section Helpers. rewrite dom_op dom_singleton_L ?gset_op_union (comm union); iFrame. iPureIntro. { by apply mark_strict_subgraph. } - + iModIntro. iLeft; iSplit; trivial. iExists _; iFrame. + + iModIntro. wp_pures. iLeft; iSplit; trivial. iExists _; iFrame. iPureIntro; eapply of_graph_unmarked; eauto. Qed. @@ -244,7 +245,7 @@ Section Helpers. - intros x; rewrite elem_of_union; intuition. - intros x y; rewrite elem_of_union; intuition eauto. Qed. - Lemma front_singleton (g : graph loc) (t : gset loc) l (w : chlC) u1 u2 : + Lemma front_singleton (g : graph loc) (t : gset loc) l (w : chlO) u1 u2 : g !! l = Some (u1, u2) → match u1 with |Some l1 => l1 ∈ t | None => True end → match u2 with |Some l2 => l2 ∈ t | None => True end → @@ -259,8 +260,8 @@ Section Helpers. Lemma empty_graph_divide q : own_graph q ∅ ⊢ own_graph (q / 2) ∅ ∗ own_graph (q / 2) ∅. Proof. - setoid_replace (∅ : gmapR loc (exclR chlC)) with - (∅ ⋅ ∅ : gmapR loc (exclR chlC)) at 1 by (by rewrite ucmra_unit_left_id). + setoid_replace (∅ : gmapR loc (exclR chlO)) with + (∅ ⋅ ∅ : gmapR loc (exclR chlO)) at 1 by (by rewrite ucmra_unit_left_id). by rewrite graph_divide. Qed.