aneris.examples.ping_pong_done.proof2

From aneris Require Import tactics proofmode notation adequacy network.
From aneris.examples.ping_pong_done Require Export code.
From iris.algebra Require Import cmra agree frac .
From stdpp Require Import base.

Set Default Proof Using "Type".

Import Network.
Import String.
Import uPred.

(* Last message received *)
Inductive msg_received := NONE | PING | PONG | DONE.

Definition ppCmra : cmraT := (prodR fracR (agreeR (leibnizO msg_received))).
Definition ticketCmra : cmraT := fracR.

Class ppdG Σ := PpdG {
  pongG_inG :> inG Σ ppCmra;
  pingG_inG :> inG Σ ppCmra;
  pongG_ticket_inG :> inG Σ ticketCmra;
  pingG_ticket_inG :> inG Σ ticketCmra;
}.

Definition pingpongΣ : gFunctors := #[GFunctor ppCmra; GFunctor ticketCmra].

Instance subG_ppdΣ {Σ} : subG pingpongΣ Σ ppdG Σ.
Proof. solve_inG. Qed.

Class pingpongG Σ := MppdG {
  pong_inG :> inG Σ ppCmra;
  ping_inG :> inG Σ ppCmra;
  pong_name : gname;
  ping_name : gname;

  pong_ticket_inG :> inG Σ ticketCmra;
  ping_ticket_inG :> inG Σ ticketCmra;
  pong_ticket_name : gname;
  ping_ticket_name : gname;
}.

Section ghost_state.
  Context `{!pingpongG Σ}.

  (* Corresponds to a view of what the last message pong received is *)
  Definition ownPong (q : Qp) (s : msg_received) :=
    own (A := ppCmra) pong_name (q, to_agree s).

  (* Corresponds to a view of what the last message ping received is *)
  Definition ownPing (q : Qp) (s : msg_received) :=
    own (A := ppCmra) ping_name (q, to_agree s).

  Lemma ownPong_eq p q s1 s2:
    ownPong p s1 -∗ ownPong q s2 -∗ s1 = s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPong.
    iCombine "H1" "H2" as "H".
    iDestruct (own_valid with "H") as %HValid.
    destruct HValid as [_ H2]; simpl in H2;
    iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
  Qed.

  Lemma ownPong_update s1 s2:
    ownPong (1/2) s1 -∗ ownPong (1/2) s1 ==∗ ownPong (1/2) s2 ownPong (1/2) s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPong. iCombine "H1 H2" as "H".
    rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
    iApply (own_update with "H"). by apply cmra_update_exclusive.
  Qed.

  Lemma ownPing_update s1 s2:
    ownPing (1/2) s1 -∗ ownPing (1/2) s1 ==∗ ownPing (1/2) s2 ownPing (1/2) s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPing. iCombine "H1 H2" as "H".
    rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
    iApply (own_update with "H"). by apply cmra_update_exclusive.
  Qed.

End ghost_state.

Section pong.
  Context `{ppd : !pingpongG Σ}
          `{dG : !distG Σ}
          `{node : Network.node}.

  (* C.f. issue in proof.v we introduce a 'Received' that covers the fact that
     the message 's' was received and send from 'from' to 'to' *)

  Definition Received (s : string) (from to : socket_address): iProp Σ :=
    ( i m, i m↦{¾} m m_body m = s m_sender m = from m_destination m = to)%I.

  Lemma mk_Received s from to i m :
    m_body m = s m_sender m = from m_destination m = to
    i m↦{¾} m -∗
      Received s from to.
  Proof.
    iIntros ((? & ? & ?)) "H".
    iExists _, _. by iFrame.
  Qed.

  Definition pong_si pong_addr : socket_interp Σ :=
    (λ msg, ϕ, ms_sender msg ϕ
     ((ms_body msg = "PING" ownPong (1/2) NONE
        ( m, (ms_body m = "PONG"
                 ms_sender m = pong_addr
                 ownPong (1/2) PING ownPing (1/2) NONE
                 Received "PING" (ms_sender msg) pong_addr) -∗ ϕ m))
       (ms_body msg = "DONE"
          ownPong (1/2) PING ownPing (1/2) PONG
          Received "PONG" pong_addr (ms_sender msg)))
    )%I.

  Arguments pong_si : simpl never.

  Lemma pong_spec (A : gset socket_address) (addr : socket_address) :
    addr A ->
    {{{ addr (pong_si addr)
         Fixed A
         FreePorts (ip_of_address addr) {[port_of_address addr]}
         IsNode node
         ownPong (1/2) NONE
         ownPing (1/2) NONE }}}
      node; pong #addr
    {{{ v, RET node;v; a, Received "PONG" addr a Received "DONE" a addr }}}.
  Proof.
    iIntros (Haddr ) "(#Hsi & #Hfixed & Hip & Hn & Hpong & Hping) Hcont".
    wp_lam.
    wp_socket h as "Hsocket". wp_let.
    wp_apply (wp_socketbind_static with "[$]"); try done.
    iDestruct 1 as (g) "(Hsocket & Hbind & _) /=". wp_pures.
    wp_bind (listen_wait _)%E. iApply (listen_wait_spec with "[$]"); try done.
    iNext. iIntros (m1 m1Id) "(Hsocket & Hsoup & Hm1 & [Heq' _] & Hsi')".
    iDestruct "Hsi'" as (ψ) "[Hsi' [[Heq (Hpong' & Hpsi)] | (_ & Hpong' & Hping' & _)]] /="; last first.
    - (* Received a DONE, which contradicts the protocol at this point *)
      iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr. inversion Hcontr.
    - (* Received a PING from somone *)
      iPoseProof (mk_Received with "Hm1") as "Hrecm1"; auto.
      iDestruct "Heq" as %->. iDestruct "Heq'" as %->.
      iDestruct (ownPong_update _ PING with "Hpong' Hpong") as ">[Hpong' Hpong]".
      wp_pures. wp_apply (wp_send_to_bound
                  with "[Hrecm1 Hsi $Hsocket $Hsi' Hpsi Hpong' Hping]"); first done.
      + (* We prove that the message is well-formed and transfer ownership of PONG's state *)
         iApply "Hpsi". by iFrame.
      + (* Sucessfully sent PONG *)
        iIntros "[Hsocket _]". wp_pures. wp_bind (listen_wait _)%E.
        iApply (listen_wait_spec with "[$]"); try done.
        iNext. iIntros (m2 m2Id) "(Hsocket & Hsoup & Hm2 & [% _] & Hsi')".
        iDestruct "Hsi'" as (ψ') "[Hsi' [[% [Hpong' Hpsi']] | (Heq & Hpong' & Hping' & Hrecv)]] /=".
        * (* Received PING, which contradicts the protocol at this point *)
          iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr. inversion Hcontr.
        * (* DONE *)
          iDestruct (ownPong_update _ DONE with "Hpong' Hpong") as ">Hpong".
          rewrite /ownPong -own_op -pair_op frac_op' Qp_half_half agree_idemp. rewrite -/(ownPong _ _).
          iPoseProof (mk_Received with "Hm2") as "Hrecm2"; try done.
          iDestruct "Heq" as %->. wp_pures.
          iApply "Hcont". by (iExists _; iFrame).
  Qed.

End pong.

Section ping.
  Context `{ppd : !pingpongG Σ}
          `{dG : distG Σ}
          `{node : Network.node}.

  Definition ping_si pong_addr ping_addr : socket_interp Σ :=
    (λ msg, ms_body msg = "PONG" ms_sender msg = pong_addr
             ownPong (1/2) PING ownPing (1/2) NONE Received "PING" ping_addr pong_addr)%I.

  Lemma ping_spec (A : gset socket_address) (addr pong : socket_address):
    pong A ->
    addr A ->
    {{{ pong (pong_si pong)
         Fixed A
         FreePorts (ip_of_address addr) {[port_of_address addr]}
         IsNode node
         ownPong (1/2) NONE
         ownPing (1/2) NONE }}}
      node; ping #addr #pong
    {{{ v, RET node; v; Received "PING" addr pong }}}.
  Proof.
    iIntros (Hserver Haddr Φ) "(#Hsi & #Hfixed & Hip & Hn & Hpong & Hping) Hcont".
    wp_lam. wp_let.
    wp_socket h as "Hsocket". wp_let.
    wp_apply (wp_socketbind_dynamic _ _ _ _ _ _ _ (ping_si pong addr) with "[$]"); try done.
    iIntros "H //=". iDestruct "H" as (g) "(Hsocket & Hbind & #Haddr)". wp_pures.
    (* Sending "PING" *)
    wp_apply (wp_send_to_bound with "[$Hsi $Hsocket Hpong]"); first done.
    - (* The message adheres to the socket interpretation *)
      iExists (ping_si pong addr). simpl. iFrame "#". iLeft.
      iSplit; first done. iFrame. iNext. iIntros (m) "[Hbody Hsipred]". by iFrame.
    - (* Sucesfully sent PING *)
      iIntros "[Hsocket Hsend]".
      wp_pures. wp_bind (listen_wait _)%E.
      iApply (listen_wait_spec with "[$]"); try done. iNext.
      iIntros (m mId) "(Hsocket & Hbind & Hm & [Hmeta _] & Heq & Heq' & Hpong' & Hping' & Hrecv) /=".
      (* Received PONG *)
      iPoseProof (mk_Received with "Hm") as "Hrecm"; auto.
      iDestruct "Heq" as %->. iDestruct "Heq'" as %->. iDestruct "Hmeta" as %->.
      iDestruct (ownPing_update _ PONG with "Hping' Hping") as ">[Hping' Hping]".
      wp_pures.
      wp_apply (wp_send_to_bound with "[$Hsi $Hsocket Hpong' Hping' Hrecm]"); first done.
      + rewrite /pong_si /=. iExists (ping_si pong addr). iFrame "#".
        iRight. iSplit; first done; iFrame.
      + iIntros "?". by iApply "Hcont".
  Qed.
End ping.

Section ping_pong_runner.
  Context `{dG : distG Σ}
          `{!pingpongG Σ}.

  Definition ips : gset string := {[ "0.0.0.0" ; "0.0.0.1"]}.
  Definition pong_addr := SocketAddressInet "0.0.0.0" 80.
  Definition ping_addr := SocketAddressInet "0.0.0.1" 80.

  Lemma ping_pong_runner_spec (A : gset socket_address) :
    pong_addr A ->
    ping_addr A ->
    {{{ pong_addr (pong_si pong_addr)
         Fixed A
         ([∗ set] ip ips, FreeIP ip)
         ownPong 1 NONE
         ownPing 1 NONE
    }}}
      ping_pong_runner
    {{{ v, RET v; True }}}.
  Proof.
    iIntros (server client Φ) "(#Hsi & #Hfix & Hips & HownPo & HownPi) Hcont".
    rewrite /ping_pong_runner.
    iDestruct (big_sepS_delete _ _ "0.0.0.0" with "Hips") as "(Hpong & Hips)";
      first set_solver.
    iDestruct (big_sepS_delete _ _ "0.0.0.1" with "Hips") as "(Hping & Hips)";
      first set_solver.
    wp_pures.
    rewrite /ownPong /ownPing -Qp_half_half -frac_op'
            -(agree_idemp (@to_agree (leibnizO _) _))
            (pair_op _ _ (@to_agree (leibnizO _ ) _) _ ) !own_op.
    iDestruct "HownPo" as "[Hpong1 Hpong2]".
    iDestruct "HownPi" as "[Hping1 Hping2]".
    (* start the nodes *)
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitR "Hpong1 Hping1"; last first.
    - iNext. iIntros "Hγs Hn".
      iApply (pong_spec _ _ server with "[$]").
      iNext. by iIntros (?) "_".
    - iNext. wp_seq.
      wp_apply (wp_start with "[-]"); first done. iFrame. iSplitL "Hcont".
      + by iApply "Hcont".
      + iNext. iIntros "Hγs Hn". iApply wp_wand_r.
        iSplitL.
        * iApply (ping_spec _ _ _ server client
                   (λ _, (Received "PING" ping_addr pong_addr)%I) with "[$]"); auto.
        * by iIntros (v) "_".
  Qed.

End ping_pong_runner.