aneris.examples.bag.bag_proof

From iris.algebra Require Import cmra agree frac .
From aneris Require Import tactics proofmode notation network.
From aneris.examples.bag Require Export bag.
From aneris.examples.library Require Export bag.

Import Network.
Import bag.

Section bag_service.
  Context `{!distG Σ, !lockG Σ}.
  Context `{node: Network.node}.
  Context (N : namespace).

  Definition inv_socket_N := N .@ "bag_socket".
  Definition inv_socket n h a := ( s g, saddress s = Some a
                                       h s↦[n] s
                                       a r↦{½} g)%I.
  Definition socket_I n h a := inv inv_socket_N (inv_socket n h a).

  Definition bag_si (Ψ : string iProp Σ) : socket_interp Σ :=
    (λ m, Φ, ms_sender m Φ
      (* inserting something into the bag satisfying Ψ *)
      ((ms_body m "" Ψ (ms_body m)
          (* the inserted element is returned as a response *)
           m', ms_body m' = "" -∗ Φ m')
      (* removing an element from the bag *)
      (ms_body m = ""
         (* either empty or an element satisfying Ψ is returned *)
          m', (ms_body m' = "" Ψ (ms_body m') -∗ Φ m'))))%I.

  Lemma serve_spec (Ψ : string iProp Σ) h a bagv :
    {{{ socket_I node h a
         isBag N node bagv (λ v, s, v = #(LitString s) Ψ s)
         a bag_si Ψ }}}
      node; serve #(LitSocket h) bagv
    {{{ v, RET node;v; False }}}.
  Proof.
    iIntros (φ) "(#Hsock & #Hbag & #Hsi) Hφ". wp_lam; wp_let. wp_closure.
    iLöb as "IH". wp_pures.
    wp_bind (ReceiveFrom _). wp_apply wp_atomic.
    iInv inv_socket_N as ">H'" "Hclose".
    iDestruct "H'" as (z g ?) "[Hz Hms]".
    iApply (wp_receive_from_2 with "[$Hz $Hms $Hsi]"); first done.
    iModIntro. iNext. iIntros (r). iDestruct 1 as "(Hz & [H | [-> Hms]])"; simpl.
    - iDestruct "H" as (m j) "(% & _ & Hms & Hj & Hsipred) /=". simplify_eq.
      iMod ("Hclose" with "[Hz Hms]") as "_".
      { iExists _,_. iFrame. eauto. }
      iModIntro. wp_pures.
      iDestruct "Hsipred" as (Φ) "[#Hsisend [(% & HΨ & HΦ) | (Heq & HΦ)]] /=".
      + case_bool_decide; wp_if; first by simplify_eq.
        wp_bind (insert _ _).
        wp_apply (bag_insert_spec with "[$Hbag HΨ]").
        { iExists _. by iFrame. }
        iIntros (?) "_ /=". wp_pures.
        wp_bind (SendTo _ _ _). wp_apply wp_atomic.
        iInv inv_socket_N as ">H'" "Hclose".
        iDestruct "H'" as (s' g' ?) "[Hz Hms]".
        iApply (wp_send_to_bound with "[$Hz $Hsisend HΦ]"); first done.
        { by iApply "HΦ". }
        iModIntro. iNext. iDestruct 1 as "[Hz _] /=".
        iMod ("Hclose" with "[Hz Hms]") as "_".
        { iExists _,_. iFrame. eauto. }
        iModIntro. wp_seq. by iApply "IH".
      + iDestruct "Heq" as %->. wp_pures. wp_bind (remove _).
        wp_apply (bag_remove_spec with "[$Hbag]").
        iIntros (?) "[-> | Helem]"; wp_pures.
        * wp_bind (SendTo _ _ _). wp_apply wp_atomic.
          iInv inv_socket_N as ">H'" "Hclose".
          iDestruct "H'" as (s' g' ?) "[Hz Hms]".
          iApply (wp_send_to_bound with "[$Hz $Hsisend HΦ]"); first done.
          { by (iApply "HΦ"; iLeft). }
          iModIntro. iNext. iDestruct 1 as "[Hz _] /=".
          iMod ("Hclose" with "[Hz Hms]") as "_".
          { iExists _,_. iFrame. eauto. }
          iModIntro. wp_seq. by iApply "IH".
        * iDestruct "Helem" as (v) "[-> HΨ]". wp_pures.
          wp_bind (SendTo _ _ _). wp_apply wp_atomic.
          iInv inv_socket_N as ">H'" "Hclose".
          iDestruct "H'" as (z'' g' ?) "[Hz Hms]". iModIntro.
          iDestruct "HΨ" as (s) "[-> HΨ]".
          iApply (wp_send_to_bound with "[$Hz $Hsisend HΨ HΦ]"); first done.
          { by (iApply "HΦ"; iRight). }
          iModIntro. iDestruct 1 as "[Hz _] /=".
          iMod ("Hclose" with "[Hz Hms]") as "_".
          { iExists _,_. iFrame. eauto. }
          iModIntro. wp_seq. by iApply "IH".
    - iMod ("Hclose" with "[Hz Hms]") as "_".
      { iExists _,_. iFrame. eauto. }
      iModIntro. wp_match. wp_seq. by iApply "IH".
  Qed.

  Lemma bag_service_spec Ψ A a :
    a A
    {{{ a bag_si Ψ
         Fixed A
         FreePorts (ip_of_address a) {[port_of_address a]}
         IsNode node }}}
      node; bag_service #a
    {{{ v, RET node;v; True }}}.
  Proof.
    iIntros (Ha Φ) "(#Hsi & #Hfa & Hp & Hnode) HΦ".
    wp_lam. wp_socket h as "Hz". wp_pures.
    wp_bind (newbag _).
    iApply (newbag_spec N _ (λ v, ( s, v = #(LitString s) Ψ s)%I) with "[$Hnode]").
    iNext. iIntros (bag) "#Hbag /=". wp_pures.
    wp_apply (wp_socketbind_static with "[$]"); try done.
    iDestruct 1 as (g) "(Hz & Hms & _) /=".
    iApply fupd_wp.
    iMod (inv_alloc inv_socket_N _ (inv_socket node h a) with "[Hz Hms]") as "#Hinv".
    { iNext. iExists _, g. by iFrame. }
    iModIntro. wp_seq.
    wp_apply wp_fork. iSplitL.
    - iNext. wp_seq.
      wp_apply wp_fork. iSplitL; first by iApply "HΦ".
      iNext. by iApply (serve_spec with "[$]").
    - by iApply (serve_spec with "[$]").
  Qed.

  Definition bag_client_si (Ψ : string iProp Σ) : socket_interp Σ :=
    (λ m, ms_body m = "" Ψ (ms_body m))%I.

  Lemma bag_client_spec Ψ A arg a server :
    server A
    a A
    {{{ server bag_si Ψ
         Fixed A
         FreePorts (ip_of_address a) {[port_of_address a]}
         IsNode node
         (arg = "" (arg "" Ψ arg))
    }}}
      node; client #arg #a #server
    {{{ v, RET node;v; v = #"" ( s, v = #(LitString s) Ψ s) }}}.
  Proof.
    iIntros (Hserver Ha Φ) "(#Hsi & #Hfa & Hp & Hnode & HΨ) HΦ".
    wp_lam. wp_socket h as "Hz". wp_pures.
    wp_apply (wp_socketbind_dynamic _ _ _ _ _ _ _ (bag_client_si Ψ) with "[$]"); try done.
    iIntros "H //=". iDestruct "H" as (g) "(Hz & ? & #Hclientsi)". wp_seq.
    wp_bind (SendTo _ _ _).
    iApply (wp_send_to_bound with "[$Hz $Hsi HΨ]"); first done.
    { rewrite /bag_si; simpl.
      iExists _. iFrame "#".
      iDestruct "HΨ" as "[% | [% HΨ]]".
      - iRight. iSplitR; first done.
        iIntros (m) "[% | HΨ]"; by [iLeft | iRight].
      - iLeft. iFrame. iSplit; try done. iIntros (m) "%". by iLeft. }
    iNext. iIntros "[Hz _] /=". wp_seq.
    wp_bind (listen_wait _).
    iApply (listen_wait_spec with "[$]"); try done.
    iNext. iIntros (m i) "(Hz & Hms & _ & _ & Hsipred) /=".
    wp_pures. iApply "HΦ". rewrite /bag_client_si. iSimplifyEq.
    iDestruct "Hsipred" as "[-> | HΨ]"; first by iLeft.
    iRight. iExists _. by iFrame.
  Qed.

End bag_service.

Section bag_runner.
  Context `{!distG Σ, !lockG Σ}.
  Context (N : namespace).

  Definition Ψeven (s : string) : iProp Σ :=
    ( (n : nat), StringOfZ n = s even n)%I.

  Definition ips : gset string := {[ "1.1.1.1" ; "0.0.0.0" ; "0.0.0.1" ; "0.0.0.2" ]}.
  Definition server := SocketAddressInet "1.1.1.1" 80.
  Definition c0 := SocketAddressInet "0.0.0.0" 80.
  Definition c1 := SocketAddressInet "0.0.0.1" 80.
  Definition c2 := SocketAddressInet "0.0.0.2" 80.

  Lemma bag_runner_spec (A : gset socket_address) :
    server A ->
    c0 A ->
    c1 A ->
    c2 A ->
    {{{ server bag_si Ψeven
         Fixed A
         ([∗ set] ip ips, FreeIP ip) }}}
      bag_runner
    {{{ v, RET v; True }}}.
  Proof.
    iIntros (???? Φ) "(#Hsi & #Hfa & Hips) HΦ".
    iDestruct (big_sepS_delete _ _ "1.1.1.1" with "Hips") as "(Hs & Hips)";
      first set_solver.
    iDestruct (big_sepS_delete _ _ "0.0.0.0" with "Hips") as "(Hc0 & Hips)";
      first set_solver.
    iDestruct (big_sepS_delete _ _ "0.0.0.1" with "Hips") as "(c1 & Hips)";
      first set_solver.
    iDestruct (big_sepS_delete _ _ "0.0.0.2" with "Hips") as "(c2 & Hips)";
      first set_solver.
    rewrite /bag_runner. wp_pures.
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitL; last first.
    { iNext. iIntros "Hnode Hp".
      iApply (bag_service_spec with "[$Hsi $Hfa $Hp $Hnode]"); first done; eauto. }
    iNext. wp_seq.
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitL; last first.
    { iNext. iIntros "Hnode Hp".
      iApply (bag_client_spec with "[$Hsi $Hfa $Hnode Hp]"); auto.
      iFrame. iRight. iSplit; first done. by iExists 2%nat. }
    iNext. wp_seq.
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitL; last first.
    { iNext. iIntros "Hnode Hp".
      iApply (bag_client_spec with "[$Hsi $Hfa $Hnode Hp]"); auto.
    }
    iNext. wp_seq.
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitL; last first.
    { iNext. iIntros "Hnode Hp".
      iApply (bag_client_spec with "[$Hsi $Hfa $Hnode Hp]"); auto.
      iFrame. iRight. iSplit; first done. by iExists 4%nat. }
    by iApply "HΦ".
  Qed.
End bag_runner.