Library iris.heap_lang.lib.increment

From iris.bi.lib Require Import fractional.
From iris.proofmode Require Import proofmode.
From iris.base_logic.lib Require Export invariants.
From iris.program_logic Require Export atomic.
From iris.heap_lang Require Import proofmode notation atomic_heap par.
From iris.prelude Require Import options.

Show that implementing fetch-and-add on top of CAS preserves logical atomicity.
First: logically atomic increment directly on top of the physical heap.

Section increment_physical.
  Context `{!heapGS Σ}.

  Definition incr_phy : val :=
    rec: "incr" "l" :=
       let: "oldv" := !"l" in
       if: CAS "l" "oldv" ("oldv" + #1)
         then "oldv"
         else "incr" "l".

  Lemma incr_phy_spec (l: loc) :
     <<{ ∀∀ (v : Z), l #v }>> incr_phy #l @ <<{ l #(v + 1) | RET #v }>>.
  Proof.
    iIntros (Φ) "AU". iLöb as "IH". wp_lam.
    wp_bind (!_)%E. iMod "AU" as (v) "[Hl [Hclose _]]".
    wp_load. iMod ("Hclose" with "Hl") as "AU".
    iModIntro. wp_pures.
    wp_bind (CmpXchg _ _ _)%E. iMod "AU" as (w) "[Hl Hclose]".
    destruct (decide (#v = #w)) as [[= ->]|Hx].
    - wp_cmpxchg_suc. iDestruct "Hclose" as "[_ Hclose]". iMod ("Hclose" with "Hl") as "HΦ".
      iModIntro. wp_pures. done.
    - wp_cmpxchg_fail. iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "Hl") as "AU".
      iModIntro. wp_pures. iApply "IH". done.
  Qed.
End increment_physical.

Next: logically atomic increment on top of an arbitrary logically atomic heap

Section increment.
  Context `{!atomic_heap}.

  Import atomic_heap.notation.

  Definition incr : val :=
    rec: "incr" "l" :=
       let: "oldv" := !"l" in
       if: CAS "l" "oldv" ("oldv" + #1)
         then "oldv"
         else "incr" "l".

  Context `{!heapGS Σ, !atomic_heapGS Σ}.

A proof of the incr specification that unfolds the definition of atomic accessors. This is the style that most logically atomic proofs take.
  Lemma incr_spec_direct (l: loc) :
     <<{ ∀∀ (v : Z), l #v }>> incr #l @ <<{ l #(v + 1) | RET #v }>>.
  Proof.
    iIntros (Φ) "AU". iLöb as "IH". wp_lam.
    awp_apply load_spec.
    rewrite /atomic_acc /=. iMod "AU" as (v) "[Hl [Hclose _]]".
    iModIntro. iExists _, _. iFrame "Hl". iSplit.
    { done. }
    iIntros "Hl". iMod ("Hclose" with "Hl") as "AU". iModIntro.
    wp_pures. awp_apply cas_spec; first done.
    rewrite /atomic_acc /=. iMod "AU" as (w) "[Hl Hclose]".
    iModIntro. iExists _. iFrame "Hl". iSplit.
    { iDestruct "Hclose" as "[? _]". done. }
    
    iIntros "Hl". simpl. destruct (decide (#w = #v)) as [[= ->]|Hx].
    - iDestruct "Hclose" as "[_ Hclose]". iMod ("Hclose" with "Hl") as "HΦ".
      iIntros "!>". wp_if. by iApply "HΦ".
    - iDestruct "Hclose" as "[Hclose _]". iMod ("Hclose" with "Hl") as "AU".
      iIntros "!>". wp_if. iApply "IH". done.
  Qed.

A proof of the incr specification that uses lemmas (aacc_aupd_×) to avoid reasoning with the definition of atomic accessors. These lemmas are only usable here because the atomic update we have and the one we try to prove are in 1:1 correspondence; most logically atomic proofs will not be able to use them.
  Lemma incr_spec (l: loc) :
     <<{ ∀∀ (v : Z), l #v }>> incr #l @ <<{ l #(v + 1) | RET #v }>>.
  Proof.
    iIntros (Φ) "AU". iLöb as "IH". wp_lam.
    awp_apply load_spec.
    iApply (aacc_aupd_abort with "AU"); first done.
    iIntros (x) "H↦". iAaccIntro with "H↦"; first by eauto with iFrame.
    iIntros "$ !> AU !>".
    wp_pures. awp_apply cas_spec; first done.
    iApply (aacc_aupd with "AU"); first done.
    iIntros (x') "H↦". iAaccIntro with "H↦"; first by eauto with iFrame.
    iIntros "H↦ !>".
    simpl. destruct (decide (#x' = #x)) as [[= ->]|Hx].
    - iRight. iFrame. iIntros "HΦ !>".
      wp_if. by iApply "HΦ".
    - iLeft. iFrame. iIntros "AU !>".
      wp_if. iApply "IH". done.
  Qed.

A "weak increment": assumes that there is no race
  Definition weak_incr: val :=
    rec: "weak_incr" "l" :=
       let: "oldv" := !"l" in
       "l" <- ("oldv" + #1);;
       "oldv" .

Logically atomic spec for weak increment. Also an example for what TaDA calls "private precondition".
  Lemma weak_incr_spec (l: loc) (v : Z) :
    l {#1/2} #v -∗
    <<{ ∀∀ (v' : Z), l {#1/2} #v' }>>
      weak_incr #l @
    <<{ v = v' l #(v + 1)
      | RET #v }>>.
  Proof.
    iIntros "Hl" (Φ) "AU". wp_lam.
    wp_apply (atomic_wp_seq $! (load_spec _) with "Hl") as "Hl".
    wp_pures. awp_apply store_spec.
    iApply (aacc_aupd_commit with "AU"); first done.
    iIntros (x) "H↦".
    iDestruct (pointsto_agree with "Hl H↦") as %[= <-].
    iCombine "Hl" "H↦" as "Hl". iAaccIntro with "Hl".
    { iIntros "[$ $]"; eauto. }
    iIntros "$ !>". iSplit; first done.
    iIntros "HΦ !>". wp_seq. done.
  Qed.

End increment.

Section increment_client.
  Context `{!heapGS Σ, !spawnG Σ}.

  Local Existing Instance primitive_atomic_heap.

  Definition incr_client : val :=
    λ: "x",
       let: "l" := ref "x" in
       incr "l" ||| incr "l".

  Lemma incr_client_safe (x: Z):
     WP incr_client #x {{ _, True }}.
  Proof using Type×.
    wp_lam. wp_alloc l as "Hl".
    iMod (inv_alloc nroot _ (x':Z, l #x')%I with "[Hl]") as "#Hinv"; first eauto.
    iAssert ( WP incr #l {{ _, True }})%I as "#Aupd".
    { iIntros "!>". awp_apply incr_spec. clear x.
      iInv nroot as (x) ">H↦". iAaccIntro with "H↦"; first by eauto 10.
      iIntros "H↦ !>". iSplitL "H↦"; first by eauto 10.
      done.
    }
    wp_smart_apply wp_par.
    - iAssumption.
    - iAssumption.
    - iIntros (??) "_ !>". done.
  Qed.

End increment_client.