LogrelCC.examples.refinement.list_assoc_binary_spec

From iris.algebra Require Import agree auth.
From iris.proofmode Require Import tactics.
From LogrelCC Require Export fundamental_binary logrel_binary rules_unary
     rules_binary.
From LogrelCC.examples Require Import list_rev
     list_assoc list_basics lock lock_unary_spec.

Section list_assoc_reltedness.
  Context `{heapG Σ, cfgSG Σ}.
  Context `{inG Σ (exclR unitR)}. (* for the locks *)

  Definition assoc_invariant γ lock lock' lst lst' f :=
    is_lock γ lock
            (lock' ↦ₛ (#♭v false)
                    ( ls : list (val * val),
                       lst ↦ᵢ embed_list (map fst ls)
                        lst' ↦ₛ embed_list (map snd ls)
                        ([∗ list] x ls, f x)))%I.

  Lemma associate_rel E ρ j K γ lock lock' lst lst' f X Y x y :
    nclose specN E
    nclose (lockN lock) E
    to_val X = Some x to_val Y = Some y
    spec_ctx ρ j fill K (App (associate (Loc lock') (Loc lst')) Y)
    assoc_invariant γ lock lock' lst lst' f f (x, y)
     CLWP App (associate (Loc lock) (Loc lst)) X @ E
    {{v, n, v = #nv n j fill K (#n n)}}.
  Proof.
    iIntros (HE HE' <-%of_to_val <-%of_to_val) "(#Hρ & Hj & #Hinv & Hf)".
    rewrite !associate_eq.
    asimpl.
    iApply clwp_Lam; eauto. iNext.
    simpl. asimpl.
    iMod (step_Lam with "[$Hj]") as "Hj"; eauto.
    simpl. asimpl.
    iApply (clwp_bind [SeqCtx _]).
    iApply (clwp_wand_r); iSplitR; first by iApply clwp_acquire; eauto.
    simpl. iIntros (v) "[[Hlc' Hls] Hil]".
    iDestruct "Hls" as (ls) "(Hlst & Hlst' & Hls)".
    iMod (steps_acquire _ _ _ (SeqCtx _ :: _) with "[$Hlc' $Hj]")
      as "[Hj Hlc']"; eauto.
    simpl.
    iApply clwp_Seq; eauto. iNext.
    simpl. asimpl.
    iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
    simpl. asimpl.
    iApply (clwp_bind [LetInCtx _]).
    iApply (clwp_load with "[-]"); iFrame; eauto.
    iNext. iIntros "Hlst". iApply clwp_value; eauto.
    simpl.
    iMod (step_load _ _ _ (LetInCtx _ :: _) with "[$Hlst' $Hj]")
      as "[Hj Hlst']"; eauto. simpl.
    iApply clwp_LetIn; eauto. iNext.
    simpl. asimpl.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto using to_of_val.
    simpl. asimpl.
    iApply (clwp_bind [LetInCtx _]).
    iApply (clwp_wand_r); iSplitR; first by iApply clwp_list_length; eauto.
    iIntros (w ?); iSimplifyEq.
    iMod (steps_list_length _ _ _ (LetInCtx _ :: _) with "[$Hj]")
      as "Hj"; eauto. simpl.
    rewrite !map_length.
    iApply clwp_LetIn; eauto. iNext.
    simpl. asimpl.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto using to_of_val.
    simpl. asimpl.
    iApply (clwp_bind [StoreRCtx (LocV _); LetInCtx _]). simpl.
    iApply (clwp_wand_r); iSplitR; first by iApply clwp_list_snoc; eauto.
    iIntros (w ?); iSimplifyEq.
    iMod (step_list_snoc _ _ _ (StoreRCtx (LocV _) :: LetInCtx _ :: _)
            with "[$Hj]") as "Hj"; eauto. simpl.
    iApply (clwp_bind [LetInCtx _]).
    iApply clwp_store; eauto using to_of_val; iFrame.
    iNext. iIntros "Hlst".
    iApply clwp_value; eauto.
    iMod (step_store _ _ _ (LetInCtx _ :: _) with "[$Hlst' $Hj]") as "[Hj Hlst']";
      eauto. simpl.
    iApply clwp_LetIn; eauto. iNext. simpl. asimpl.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
    iApply (clwp_bind [SeqCtx _]).
    iMod (steps_release _ _ _ (SeqCtx _ :: _) with "[$Hlc' $Hj]")
      as "[Hj Hlc']"; eauto. simpl.
    iApply (clwp_wand_r); iSplitR "Hj".
    - iApply clwp_release; eauto.
      iFrame "#"; iFrame.
      iExists (ls ++ [(x, y)]).
      by rewrite !map_app /=; iFrame.
    - iIntros (?) "_".
      iApply clwp_Seq; eauto. iNext. simpl.
      iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
      iApply clwp_value; eauto.
  Qed.

  Lemma get_rel E ρ j K γ lock lock' lst lst' f n :
    ( x, Persistent (f x))
    nclose specN E
    nclose (lockN lock) E
    spec_ctx ρ j fill K (App (get (Loc lock') (Loc lst')) (#n n))
    assoc_invariant γ lock lock' lst lst' f
     CLWP App (get (Loc lock) (Loc lst)) (#n n) @E
    {{v, (v = InjLV UnitV j fill K (InjL Unit))
          w w', f (w, w') v = InjRV w j fill K (InjR (of_val w'))}}.
  Proof.
    iIntros (Hf HE HE') "(#Hρ & Hj & #Hinv)".
    rewrite !get_eq.
    asimpl.
    iApply clwp_Lam; eauto. iNext.
    simpl. asimpl.
    iApply (clwp_bind [SeqCtx _]).
    iApply (clwp_wand_r); iSplitR; first by iApply clwp_acquire; eauto.
    simpl. iIntros (v) "[[Hlc' Hls] Hil]".
    iDestruct "Hls" as (ls) "(Hlst & Hlst' & Hls)".
    iMod (step_Lam with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
    iMod (steps_acquire _ _ _ (SeqCtx _ :: _) with "[$Hlc' $Hj]")
      as "[Hj Hlc']"; eauto.
    simpl.
    iApply clwp_Seq; eauto. iNext.
    iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
    iApply (clwp_bind [LetInCtx _]).
    iApply (clwp_load with "[-]"); iFrame; eauto.
    iNext. iIntros "Hlst". iApply clwp_value; eauto. simpl.
    iMod (step_load _ _ _ (LetInCtx _ :: _)
            with "[$Hlst' $Hj]") as "[Hj Hlst']"; eauto. simpl.
    iApply clwp_LetIn; eauto. iNext.
    simpl. asimpl.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
    simpl. asimpl.
    iApply (clwp_bind [LetInCtx _]).
    iApply (clwp_wand_r); iSplitR; first by iApply clwp_list_get; eauto.
    rewrite list_lookup_fmap. iIntros (w Hw).
    iMod (step_list_get _ _ _ (LetInCtx _ :: _) with "[$Hj]")
      as (w') "[Hj Hw']"; eauto. simpl. iDestruct "Hw'" as %Hw'.
    rewrite list_lookup_fmap in Hw'.
    iApply clwp_LetIn; eauto. iNext.
    simpl. asimpl.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
    simpl. asimpl.
    iApply (clwp_bind [SeqCtx _]).
    iMod (steps_release _ _ _ (SeqCtx _ :: _) with "[$Hlc' $Hj]")
      as "[Hj Hlc']"; eauto. simpl.
    destruct (ls !! n) as [[? ?]|] eqn:Heq; simpl; simpl in *; subst.
    - iDestruct (big_sepL_lookup with "Hls") as "#Hf"; eauto.
      iApply (clwp_wand_r); iSplitR "Hj".
      + iApply clwp_release; eauto.
        iFrame "#"; iFrame. iExists ls; iFrame.
      + iIntros (?) "_".
        iApply clwp_Seq; eauto. iNext. simpl.
        iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
        iApply clwp_value; eauto.
        iRight; eauto.
    - iApply (clwp_wand_r); iSplitR "Hj".
      + iApply clwp_release; eauto.
        iFrame "#"; iFrame. iExists ls; iFrame.
      + iIntros (?) "_".
        iApply clwp_Seq; eauto. iNext. simpl.
        iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
        iApply clwp_value; eauto.
  Qed.

End list_assoc_reltedness.