LogrelCC.examples.refinement.stack.refinement

From iris.algebra Require Import auth.
From iris.program_logic Require Import adequacy.
From LogrelCC Require Import rules_unary rules_binary soundness_binary.
From LogrelCC.examples.refinement Require Import lock.
From LogrelCC.examples.refinement.stack Require Import
  CG_stack FG_stack stack_rules.
From iris.proofmode Require Import tactics.

Definition stackN : namespace := nroot .@ "stack".

Section Stack_refinement.
  Context `{heapG Σ, cfgSG Σ, inG Σ (authR stackUR)}.
  Notation D := (prodC valC valC -n> iProp Σ).
  Implicit Types Δ : listC D.

  Lemma FG_CG_counter_refinement :
    [] FG_stack log CG_stack : TForall (TProd (TProd
           (TArrow (TVar 0) TUnit)
           (TArrow TUnit (TSum TUnit (TVar 0))))
           (TArrow (TArrow (TVar 0) TUnit) TUnit)).
  Proof.
    (* executing the preambles *)
    iIntros (Δ [|??] ρ ?) "#[Hspec HΓ]"; iIntros (KK j) "[Hj HKK]"; last first.
    { iDestruct (interp_env_length with "HΓ") as %[=]. }
    iClear "HΓ". cbn -[FG_stack CG_stack].
    rewrite ?empty_env_subst /CG_stack /FG_stack.
    iApply ("HKK" $! (TLamV _, TLamV _)); iFrame; clear KK j.
    iAlways; iIntrosi Hτi KK j) "[Hj HKK] //=".
    iMod (step_tlam with "[Hj]") as "Hj"; eauto.
    iApply wp_tapp; iNext.
    iMod (steps_newlock _ _ j ((LetInCtx _) :: _) with "[$Hj]")
      as (l) "[Hj Hl]"; eauto.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
    simpl.
    rewrite CG_locked_push_subst CG_locked_pop_subst
            CG_iter_subst CG_snap_subst. simpl. asimpl.
    iMod (step_alloc _ _ j ((LetInCtx _) :: _) with "[Hj]")
      as (stk') "[Hj Hstk']"; [| |simpl; by iFrame|]; auto.
    iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
    simpl.
    rewrite CG_locked_push_subst CG_locked_pop_subst
            CG_iter_subst CG_snap_subst. simpl. asimpl.
    iApply (wp_alloc (FoldCtx :: AllocCtx :: (LetInCtx _) :: _)); eauto.
    iNext. iIntros (istk) "Histk /=".
    iApply (wp_alloc ((LetInCtx _) :: _)); eauto.
    iNext. iIntros (stk) "Hstk /=".
    iApply wp_LetIn; eauto; iNext; simpl.
    rewrite FG_push_subst FG_pop_subst FG_iter_subst. simpl. asimpl.
    (* establishing the invariant *)
    iMod (own_alloc ( ( : stackUR))) as (γ) "Hemp"; first done.
    set (istkG := StackG _ _ γ).
    change γ with (@stack_name _ istkG).
    change H1 with (@stack_inG _ istkG).
    clearbody istkG. clear γ H1.
    iAssert (@stack_owns _ istkG _ ) with "[Hemp]" as "Hoe".
    { rewrite /stack_owns big_sepM_empty. iFrame "Hemp"; trivial. }
    iMod (stack_owns_alloc with "[$Hoe $Histk]") as "[Hoe Hls]".
    iAssert (StackLink τi (LocV istk, FoldV (InjLV UnitV))) with "[Hls]" as "HLK".
    { rewrite StackLink_unfold.
      iExists _, _. iSplitR; simpl; trivial.
      iFrame "Hls". iLeft. iSplit; trivial. }
    iAssert (( istk v h, (stack_owns h)
                          stk' ↦ₛ v
                          stk ↦ᵢ (FoldV (LocV istk))
                          StackLink τi (LocV istk, v)
                          l ↦ₛ (#♭v false)
             )%I) with "[Hoe Hstk Hstk' HLK Hl]" as "Hinv".
    { iExists _, _, _. by iFrame "Hoe Hstk' Hstk Hl HLK". }
    iMod (inv_alloc stackN with "[Hinv]") as "#Hinv"; [iNext; iExact "Hinv"|].
    clear istk.
    Opaque stack_owns.
    (* splitting *)
    rewrite -(of_to_val _ _ (FG_push_to_val _))
    -(of_to_val _ _ (FG_pop_to_val _))
    -CG_locked_push_of_val -CG_locked_pop_of_val.
    iApply ("HKK" $! (PairV (PairV _ _) (LamV _), PairV (PairV _ _) (LamV _)));
      simpl; iFrame; clear j KK.
    iExists (_, _), (_, _); iSplit; eauto. iSplit.
    (* refinement of push and pop *)
    - iExists (_, _), (_, _); iSplit; eauto. iSplit.
      + (* refinement of push *)
        iAlways. iIntros ([v1 v2]) "#Hrel". iIntros (KK j) "[Hj #HKK] /=".
        rewrite (of_to_val _ _ (FG_push_to_val _))
                CG_locked_push_of_val.
        iLöb as "Hlat".
        rewrite {2}(FG_push_folding (Loc stk)).
        iApply wp_rec; eauto using to_of_val.
        iNext.
        rewrite -(FG_push_folding (Loc stk)).
        asimpl.
        iApply (wp_atomic_under_ectx _ _ ((LetInCtx _) :: _)); eauto.
        iInv stackN as (istk v h) "[Hoe [Hstk' [Hstk [HLK Hl]]]]" "Hclose";
          iModIntro.
        iApply (wp_load' with "[-]"); iFrame. iNext. iIntros "Hstk".
        iApply wp_value; eauto.
        iMod ("Hclose" with "[Hoe Hstk' HLK Hl Hstk]") as "_".
        { iNext. iExists _, _, _; by iFrame "Hoe Hstk' HLK Hl Hstk". }
        clear v h. iModIntro.
        iApply wp_LetIn; eauto using to_of_val.
        iNext. asimpl.
        iApply (wp_alloc (FoldCtx :: CasRCtx (LocV _) (FoldV (LocV _))
                                  :: IfCtx _ _ :: _)); auto.
        iNext. iIntros (ltmp) "Hltmp /=".
        iApply (wp_atomic_under_ectx _ _ (IfCtx _ _ :: _)); eauto.
        iInv stackN as (istk2 v h) "[Hoe [Hstk' [Hstk [HLK Hl]]]]" "Hclose".
        (* deciding whether CAS will succeed or fail *)
        destruct (decide (istk = istk2)) as [|Hneq]; subst.
        * (* CAS succeeds *)
          (* In this case, the specification pushes *)
          iMod "Hstk'". iMod "Hl".
          iMod (steps_CG_locked_push with "[Hj Hl Hstk']")
            as "[Hj [Hstk' Hl]]"; first solve_ndisj.
          { by iFrame "Hspec Hstk' Hj". }
          iModIntro.
          iApply (wp_cas_suc' with "[-]"); auto; iFrame.
          iNext. iIntros "Hstk".
          iMod (stack_owns_alloc with "[$Hoe $Hltmp]") as "[Hoe Hpt]".
          iApply wp_value; eauto.
          iMod ("Hclose" with "[-Hj]") as "_".
          { iNext. iExists ltmp, _, _.
            iFrame "Hoe Hstk' Hstk Hl".
            do 2 rewrite StackLink_unfold.
            rewrite -StackLink_unfold.
            iExists _, _. iSplit; trivial.
            iFrame "Hpt". eauto 10. }
          iModIntro; simpl.
          iApply wp_if_true. iNext.
          iApply ("HKK" $! (UnitV, UnitV)); eauto.
        * iModIntro.
          iApply (wp_cas_fail' with "[-]"); auto; try iFrame.
          { congruence. }
          iNext. iIntros "Hstk". iApply wp_value; eauto.
          iMod ("Hclose" with "[-Hj]").
          { iNext. iExists _, _, _. by iFrame "Hoe Hstk' Hstk Hl". }
          iApply wp_if_false. iModIntro. iNext. by iApply "Hlat".
      + (* refinement of pop *)
        iAlways. iIntros ( [v1 v2] ) "[% %]"; simpl in *; simplify_eq.
        iIntros (KK j) "[Hj HKK] /="; simplify_eq/=.
        rewrite (of_to_val _ _ (FG_pop_to_val _))
                CG_locked_pop_of_val.
        iLöb as "Hlat".
        rewrite {2}(FG_pop_folding (Loc stk)).
        iApply wp_rec; eauto using to_of_val.
        iNext.
        rewrite -(FG_pop_folding (Loc stk)).
        asimpl.
        iApply (wp_atomic_under_ectx _ _ (UnfoldCtx :: (LetInCtx _) :: _));
          eauto.
        iInv stackN as (istk v h) "[Hoe [Hstk' [Hstk [#HLK Hl]]]]" "Hclose".
        iApply (wp_load' with "[-]"); iFrame. iNext. iIntros "Hstk".
        iApply wp_value; eauto.
        iPoseProof "HLK" as "HLK'".
        (* Checking whether the stack is empty *)
        rewrite {2}StackLink_unfold.
        iDestruct "HLK'" as (istk2 w) "[% [Hmpt [[% %]|HLK']]]"; simplify_eq/=.
        * (* The stack is empty *)
          iMod (steps_CG_locked_pop_fail with "[$Hspec $Hstk' $Hl $Hj]")
             as "[Hj [Hstk' Hl]]"; first solve_ndisj.
          iMod ("Hclose" with "[-Hj Hmpt HKK]") as "_".
          { iNext. iExists _, _, _. by iFrame "Hoe Hstk' Hstk Hl". }
          iModIntro.
          iApply (wp_unfold ((LetInCtx _) :: _)); eauto.
          iNext. iApply wp_LetIn; eauto using to_of_val. iNext. asimpl.
          iClear (h) "HLK".
          iApply (wp_atomic_under_ectx _ _ ((LetInCtx _) :: _));
            eauto.
          iInv stackN as (istk3 w h) "[Hoe [Hstk' [Hstk [HLK Hl]]]]" "Hclose".
          iDestruct (stack_owns_later_open_close with "Hoe Hmpt") as "[Histk HLoe]".
          iModIntro. iApply (wp_load' with "[-]"); iFrame.
          iNext. iIntros "Histk". iApply wp_value; eauto.
          iMod ("Hclose" with "[-Hj HKK]") as "_".
          { iNext. iExists _, _, _. iFrame "Hstk' Hstk HLK Hl".
            by iApply "HLoe". }
          iApply wp_LetIn; simpl; eauto.
          iModIntro. iNext. asimpl.
          iApply wp_case_injl; eauto; simpl.
          iNext.
          iApply ("HKK" $! (InjLV UnitV, InjLV UnitV)); iFrame.
          iLeft; iExists (_, _); eauto.
        * (* The stack is not empty *)
          iMod ("Hclose" with "[-Hj Hmpt HLK' HKK]") as "_".
          { iNext. iExists _, _, _. by iFrame "Hstk' Hstk HLK Hl". }
          iModIntro.
          iApply (wp_unfold ((LetInCtx _) :: _)); eauto.
          iNext. iApply wp_LetIn; eauto using to_of_val.
          iNext. asimpl.
          iClear (h) "HLK".
          iApply (wp_atomic_under_ectx _ _ ((LetInCtx _) :: _)); eauto.
          iInv stackN as (istk3 w' h) "[Hoe [Hstk' [Hstk [HLK Hl]]]]" "Hclose".
          iModIntro.
          iDestruct (stack_owns_later_open_close with "Hoe Hmpt") as "[Histk HLoe]".
          iApply (wp_load' with "[-]"); iFrame. iNext; iIntros "Histk".
          iDestruct ("HLoe" with "Histk") as "Hh".
          iApply wp_value; eauto using to_of_val.
          iMod ("Hclose" with "[-Hj Hmpt HLK' HKK]") as "_".
          { iNext. iExists _, _, _. by iFrame "Hstk' Hstk HLK Hl". }
          iModIntro.
          iApply wp_LetIn; eauto using to_of_val.
          iNext. asimpl.
          iDestruct "HLK'" as (y1 z1 y2 z2) "[% HLK']". subst. simpl.
          iApply wp_case_injr; [simpl; rewrite ?to_of_val; eauto |].
          iNext.
          iApply (wp_snd (CasRCtx (LocV _) (FoldV (LocV _)) :: IfCtx _ _ :: _));
            eauto. simpl. iNext.
          clear istk3 h. asimpl.
          iApply (wp_atomic_under_ectx _ _ (IfCtx _ _ :: _)); eauto.
          iInv stackN as (istk3 w h) "[Hoe [Hstk' [Hstk [#HLK Hl]]]]" "Hclose".
          (* deciding whether CAS will succeed or fail *)
          destruct (decide (istk2 = istk3)) as [|Hneq]; subst.
          -- (* CAS succeeds *)
            (* In this case, the specification pushes *)
            iApply (wp_cas_suc' with "[-]");eauto. iFrame.
            iNext. iIntros "Hstk {HLK'}". iPoseProof "HLK" as "HLK'".
            rewrite {2}StackLink_unfold.
            iDestruct "HLK'" as (istk4 w2) "[% [Hmpt' HLK']]"; simplify_eq/=.
            iDestruct (stack_mapstos_agree with "[Hmpt Hmpt']") as %?;
              first (iSplit; [iExact "Hmpt"| iExact "Hmpt'"]).
            iDestruct "HLK'" as "[[% %]|HLK']"; simplify_eq/=.
            iDestruct "HLK'" as (yn1 yn2 zn1 zn2)
                                   "[% [% [#Hrel HLK'']]]"; simplify_eq/=.
             (* Now we have proven that specification can also pop. *)
             iMod (steps_CG_locked_pop_suc with "[$Hspec $Hstk' $Hl $Hj]")
                as "[Hj [Hstk' Hl]]"; first solve_ndisj.
             iApply wp_value; eauto.
             iMod ("Hclose" with "[-Hj HKK]") as "_".
             { iNext. iIntros "{Hmpt Hmpt' HLK}". rewrite StackLink_unfold.
               iDestruct "HLK''" as (istk5 w2) "[% [Hmpt HLK]]"; simplify_eq/=.
               iExists istk5, _, _. iFrame "Hoe Hstk' Hstk Hl".
               rewrite StackLink_unfold.
               iExists _, _; iSplitR; trivial.
               by iFrame "HLK". }
             iModIntro. iApply wp_if_true. iNext.
             iApply (wp_fst (InjRCtx :: _)); eauto. iNext.
             iApply ("HKK" $! (InjRV _, InjRV _)); iFrame.
             iRight; iExists (_, _); eauto.
          -- (* CAS will fail *)
            iModIntro.
            iApply (wp_cas_fail' with "[-]"); try iFrame; eauto; try congruence.
            iNext. iIntros "Hstk". iApply wp_value; eauto.
            iMod ("Hclose" with "[-Hj HKK]") as "_".
            { iNext. iExists _, _, _. by iFrame "Hoe Hstk' Hstk HLK Hl". }
            iModIntro.
            iApply wp_if_false. iNext. iApply ("Hlat" with "[Hj]"); eauto.
    - (* refinement of iter *)
      iAlways. iIntros ( [f1 f2] ) "/= #Hfs". iIntros (KK j) "[Hj #HKK] /=".
      iApply wp_Lam; eauto using to_of_val. iNext.
      iMod (step_Lam with "[$Hspec $Hj]") as "Hj"; eauto.
      asimpl. rewrite FG_iter_subst CG_snap_subst CG_iter_subst. asimpl.
      replace (FG_iter (of_val f1)) with (of_val (FG_iterV (of_val f1)))
        by (by rewrite FG_iter_of_val).
      replace (CG_iter (of_val f2)) with (of_val (CG_iterV (of_val f2)))
        by (by rewrite CG_iter_of_val).
      iApply (wp_atomic_under_ectx _ _ (AppRCtx _ :: _)); eauto.
      iInv stackN as (istk3 w h) "[Hoe [>Hstk' [>Hstk [#HLK >Hl]]]]" "Hclose";
        iModIntro.
      iMod (steps_CG_snap _ _ _ (AppRCtx _ :: _)
            with "[Hstk' Hj Hl]") as "[Hj [Hstk' Hl]]"; first solve_ndisj.
      { rewrite ?fill_app. simpl. by iFrame "Hspec Hstk' Hl Hj". }
      iApply (wp_load' with "[-]"); iFrame. iNext. iIntros "Hstk".
      iApply wp_value; eauto.
      iMod ("Hclose" with "[-Hj HKK]") as "_".
      { iNext. iExists _, _, _; by iFrame "Hoe Hstk' Hstk HLK Hl". }
      clear h.
      rewrite ?fill_app /= -FG_iter_folding.
      iModIntro.
      iLöb as "Hlat" forall (j istk3 w) "Hj HLK".
      rewrite {2}FG_iter_folding.
      iApply wp_rec; simpl; eauto.
      rewrite -FG_iter_folding. asimpl. rewrite FG_iter_subst.
      iNext.
      iApply (wp_unfold (LoadCtx :: CaseCtx _ _ :: _)); eauto; iNext; simpl.
      rewrite StackLink_unfold.
      iDestruct "HLK" as (istk4 v) "[% [Hmpt HLK]]"; simplify_eq/=.
      iApply (wp_atomic_under_ectx _ _ ( CaseCtx _ _ :: _)); eauto.
      iInv stackN as (istk5 v' h) "[Hoe [Hstk' [Hstk [HLK' Hl]]]]" "Hclose".
      iDestruct (stack_owns_later_open_close with "Hoe Hmpt") as "[Histk HLoe]".
      iModIntro.
      iApply (wp_load' with "[-]"); iFrame. iNext. iIntros "Histk".
      iDestruct ("HLoe" with "Histk") as "Hh".
      iDestruct "HLK" as "[[% %]|HLK'']"; simplify_eq/=.
      * rewrite CG_iter_of_val.
        iMod (steps_CG_iter_end with "[$Hspec $Hj]") as "Hj"; first solve_ndisj.
        iApply wp_value; eauto.
        iMod ("Hclose" with "[-Hj HKK]").
        { iNext. iExists _, _, _. by iFrame "Hh Hstk' Hstk Hl". }
        iApply wp_case_injl; eauto.
        iModIntro. iNext.
        iApply ("HKK" $! (UnitV, UnitV)); eauto.
      * iDestruct "HLK''" as (yn1 yn2 zn1 zn2)
                              "[% [% [#Hrel HLK'']]]"; simplify_eq/=.
        rewrite CG_iter_of_val.
        iMod (steps_CG_iter with "[$Hspec $Hj]") as "Hj"; first solve_ndisj.
        iApply wp_value; first rewrite /IntoVal /= !to_of_val //.
        iMod ("Hclose" with "[-Hj HKK HLK'']") as "_".
        { iNext. iExists _, _, _. by iFrame "Hh Hstk' Hstk Hl". }
        simpl.
        iApply wp_case_injr; eauto.
        simpl. rewrite FG_iter_subst CG_iter_subst. asimpl.
        iModIntro. iNext.
        iApply (wp_fst (AppRCtx _ :: (LetInCtx _) :: _)); eauto.
        iNext; simpl.
        rewrite StackLink_unfold.
        iDestruct "HLK''" as (istk6 w') "[% HLK]"; simplify_eq/=.
        iSpecialize ("Hfs" $! (yn1, zn1) with "Hrel").
        iApply ("Hfs" $! ((LetInCtx _) :: _, (LetInCtx _) :: _));
          iFrame. clear j.
        iAlways; iIntros ([] j) "[Hj #HKK2]". simpl.
        iMod (step_LetIn with "[$Hspec $Hj]") as "Hj"; eauto using to_of_val.
        asimpl. rewrite CG_iter_subst. asimpl.
        replace (CG_iter (of_val f2)) with (of_val (CG_iterV (of_val f2)))
          by (by rewrite CG_iter_of_val).
        iMod (step_snd _ _ _ (AppRCtx _ :: _) with "[$Hspec Hj]") as "Hj";
          [| | |simpl; by iFrame "Hj"|]; rewrite ?to_of_val; auto.
        iApply wp_LetIn; simpl; eauto using to_of_val.
        iNext; simpl. rewrite FG_iter_subst. asimpl.
        replace (FG_iter (of_val f1)) with (of_val (FG_iterV (of_val f1)))
          by (by rewrite FG_iter_of_val).
        iApply (wp_snd (AppRCtx _ :: _)); eauto using to_of_val.
        iNext;simpl. rewrite -FG_iter_folding.
        iApply ("Hlat" $! j istk6 zn2 with "[Hj] [HLK]"); trivial.
        rewrite StackLink_unfold; iAlways; simpl.
        iDestruct "HLK" as "[HLK Hei]".
        iExists _, _; iSplit; eauto; iFrame.
        iFrame "HLK".
        iDestruct "Hei" as "[Hei|Hei]"; auto.
        iRight. iDestruct "Hei" as (? ? ? ?) "(? & ? & ? & ?)".
        iExists _, _, _, _; iFrame "#".
  Qed.

End Stack_refinement.

Theorem stack_ctx_refinement :
  [] FG_stack ctx CG_stack : TForall (TProd (TProd
        (TArrow (TVar 0) TUnit)
        (TArrow TUnit (TSum TUnit (TVar 0))))
        (TArrow (TArrow (TVar 0) TUnit) TUnit)).
Proof.
  set (Σ := #[invΣ; gen_heapΣ loc val; GFunctor (authR cfgUR); GFunctor (authR stackUR)]).
  set (HG := soundness_unary.HeapPreIG Σ _ _).
  eapply (binary_soundness Σ); eauto using FG_stack_closed, CG_stack_closed.
  intros; apply FG_CG_counter_refinement.
Qed.