RunST.lam_beta.lam_beta_part3

From RunST Require Import IC base lang rules typing reduction regions
     future contextual_refinement saved_pred ST_Lang_reduction
     logrel_shared soundness logrel fundamental.
From iris.proofmode Require Import tactics.

Hint Resolve to_of_val.

Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_swap_201 : (lookup_n [2; 0; 1]) = (2 .: 0 .: 1 .: (+3)).
Proof.
  extensionality i; destruct i as [|[|[|?]]]; by cbv.
Qed.

Lemma lam_beta_push_lam_case Γ e2' τ e0 e1 e2 τ1 τ2 τ3 :
  Γ ⊢ₜ e2' : τ τ :: Γ ⊢ₜ e0 : TSum τ1 τ2
  τ1 :: τ :: Γ ⊢ₜ e1 : τ3 τ2 :: τ :: Γ ⊢ₜ e2 : τ3
  Γ App (Rec (Case e0.[ren (+1)] e1.[up (ren (+1))] e2.[up (ren (+1))])) e2'
  ctx Case (App (Rec e0.[ren (+1)]) e2')
          (App (Rec e1.[up (ren (+1))].[ren (lookup_n (swap_list 2 1))])
             e2'.[ren (+1)])
          (App (Rec e2.[up (ren (+1))].[ren (lookup_n (swap_list 2 1))])
             e2'.[ren (+1)]) : τ3.
Proof.
  intros Ht2' Ht0 Ht1 Ht2.
  apply binary_soundness.
  { repeat econstructor; eauto.
    - apply (context_weakening [_]); eauto.
    - apply (context_gen_weakening [_] [_]); eauto.
    - apply (context_gen_weakening [_] [_]); eauto. }
  { repeat econstructor; eauto.
    - apply (context_weakening [_]); eauto.
    - eapply (context_swap_typed [] [_] [_; _]).
      apply (context_gen_weakening [_] [_]); simpl; eauto.
    - apply (context_weakening [_]); eauto.
    - eapply (context_swap_typed [] [_] [_; _]).
      apply (context_gen_weakening [_] [_]); simpl; eauto.
    - apply (context_weakening [_]); eauto. }
  iIntros (Σ ST Hlr Δ vvs ) "[#Hinv #Hvvs]"; asimpl.
  iDestruct (interp_env_length with "[]") as %Hvvsl; eauto.
  iIntros (γh γh' σ1') "Hσ1' /=".
  rewrite ic_eq /ic_def.
  iIntros (σ1 σ2 vf n) "(Hrd & Hσ1)".
  iDestruct "Hrd" as %Hrd.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd.
  destruct Hrd as (k1 & σ3 & w2 & Hk1 & Hrd1 & Hrd2). simpl in *.
  apply rec_red_step in Hrd2.
  destruct Hrd2 as (m & Hm & Hrd2).
  asimpl in Hrd2.
  repeat (erewrite typed_subst_head_simpl in Hrd2; eauto;
          last by rewrite /= !fmap_length Hvvsl).
  apply (nsteps_bind (fill_item (CaseCtx _ _))) in Hrd2.
  destruct Hrd2 as (k2 & σ4 & w1 & Hk2 & Hrd21 & Hrd22). simpl in *.
  set (Q := (|={}=> Nat.iter k1 (λ Q, Q) ( ( (σ3' : state) (w2' : val),
           (rtc (pstep' det_lang)
          (e2'.[env_subst (vvs.*2)], σ1') (of_val w2', σ3')))))%I : iProp Σ).
  iAssert Q as "HQ".
  { unfold Q.
    iMod (allocate_full_state σ1) as (γh3) "[Hfσ1 Hnσ1]".
    iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
    iPoseProof (binary_fundamental _ _ _ Ht2' Δ vvs with "[]") as "Hrel";
      first iFrame "#".
    iSpecialize ("Hrel" $! _ _ _ with "[$Hfσ1']"). asimpl.
    rewrite ic_eq /ic_def.
    iSpecialize ("Hrel" $! _ _ _ _ with "[$Hfσ1]"); auto.
    iApply (future_plain with "[-]").
    iMod "Hrel" as "[Hrel _]". iModIntro.
    iDestruct "Hrel" as (σ3' ?) "(Hrd2&Hσ3'&?)".
    iDestruct "Hrd2" as %Hrd2.
    iExists _, _; iPureIntro; repeat split; eauto. }
  iMod "HQ" as "#HQ". iModIntro. clear Q.
  rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
  rewrite later_n_except_0_exist. iDestruct "HQ" as (w2') "HQ".
  set (Q := (|={}=> Nat.iter (k1 + k2) (λ Q, Q) ( ( (σ4' : state) (w1' : val),
           (rtc (pstep' det_lang)
          (e0.[env_subst (w2' :: vvs.*2)], σ3') (of_val w1', σ4')))))%I : iProp Σ).
  iAssert Q as "HQ2".
  { unfold Q.
    iMod (allocate_full_state σ1) as (γh3) "[Hfσ1 Hnσ1]".
    iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
    iPoseProof (binary_fundamental _ _ _ Ht2' Δ vvs with "[]") as "Hrel";
      first iFrame "#".
    iSpecialize ("Hrel" $! _ _ _ with "[$Hfσ1']"). asimpl.
    rewrite ic_eq /ic_def.
    iSpecialize ("Hrel" $! _ _ _ _ with "[$Hfσ1]"); auto.
    iApply (future_plain with "[-]").
    iDestruct (later_n_except_0_future with "[$HQ]") as "HQ'".
    iCombine "Hrel" "HQ'" as "Hrel".
    iMod "Hrel" as "[[Hrel Hσ3] HQ']". iDestruct "HQ'" as %HQ.
    iDestruct "Hrel" as (σ3'' w2'') "(Hrd2&Hσ3''&#Hww2)".
    iDestruct "Hrd2" as %Hrd2.
    apply rtc_nsteps in HQ. destruct HQ as [n1 HQ].
    apply rtc_nsteps in Hrd2. destruct Hrd2 as [n2 Hrd2].
    destruct (nsteps_deterministic' Hrd2 HQ) as [? [? ?]]; subst.
    iPoseProof (binary_fundamental _ _ _ Ht0 Δ ((w2, w2') :: vvs) with "[]")
      as "Hrel"; first by rewrite interp_env_cons; iFrame "#".
    iSpecialize ("Hrel" $! _ _ _ with "[$Hσ3'']"). asimpl.
    rewrite ic_eq /ic_def.
    iSpecialize ("Hrel" $! _ _ _ _ with "[$Hσ3]"); auto.
    iMod "Hrel" as "[Hrel _]".
    iDestruct "Hrel" as (σ4' w1') "(Hrd3&Hσ4'&?)".
    iDestruct "Hrd3" as %Hrd3. iModIntro.
    iExists _, _; iPureIntro; repeat split; eauto. }
  iMod "HQ2" as "#HQ2". iModIntro. clear Q.
  rewrite later_n_except_0_exist. iDestruct "HQ2" as (σ4') "HQ2".
  rewrite later_n_except_0_exist. iDestruct "HQ2" as (w1') "HQ2".
  iDestruct (later_n_except_0_future with "[$HQ]") as "HQ3".
  iDestruct (later_n_except_0_future with "[$HQ2]") as "HQ4".
  iClear "HQ HQ2".
  iPoseProof (binary_fundamental _ _ _ Ht2' Δ vvs with "[]") as "He21";
      first iFrame "#".
  iSpecialize ("He21" $! _ _ _ with "[$Hσ1']"). asimpl.
  rewrite ic_eq /ic_def.
  iSpecialize ("He21" $! _ _ _ _ with "[$Hσ1]"); auto.
  iMod (allocate_full_state σ1) as (γh3) "[Hfσ1 _]".
  iMod (allocate_full_state σ4') as (γh4) "[Hfσ4' Hnσ4']".
  iPoseProof (binary_fundamental _ _ _ Ht2' Δ vvs with "[]") as "He22";
    first iFrame "#".
  iSpecialize ("He22" $! _ _ _ with "[$Hfσ4']"). asimpl.
  rewrite ic_eq /ic_def.
  iSpecialize ("He22" $! _ _ _ _ with "[$Hfσ1]"); auto.
  iCombine "HQ3" "He21" as "He21". iCombine "He21" "He22" as "He2".
  iPoseProof (future_cancel_2 with "[He2 HQ4]") as "He"; first trivial.
  { iSplitL "HQ4"; iFrame. }
  iMod "He" as "(HQ' & [HQ [He21 Hσ3]] & [He22 Hfσ3])".
  iDestruct "HQ" as %HQ.
  apply rtc_nsteps in HQ. destruct HQ as [n1 HQ].
  iDestruct "He21" as (σ3'' w2'') "(Hrd2 & Hσ3' & #Hww2)".
  iDestruct "Hrd2" as %Hrd2.
  apply rtc_nsteps in Hrd2. destruct Hrd2 as [n2 Hrd2].
  destruct (nsteps_deterministic' Hrd2 HQ) as [? [? ?]]; subst.
  iDestruct "He22" as (σ5' w2'') "(Hrd3 & Hσ5' & #Hww2')".
  iDestruct "Hrd3" as %Hrd3.
  apply rtc_nsteps in Hrd3. destruct Hrd3 as [n3 Hrd3].
  replace (k1 + k2 - k1) with k2 by omega.
  iDestruct (heap_included with "[$Hnσ4' $Hσ5']") as %[Hincl Hvl].
  iClear "Hnσ4' Hσ5' Hfσ3".
  iPoseProof (binary_fundamental _ _ _ Ht0 Δ ((w2, w2') :: vvs) with "[]")
    as "He1"; first by rewrite interp_env_cons; iFrame "#".
  iSpecialize ("He1" $! _ _ _ with "[$Hσ3']"). asimpl.
  rewrite ic_eq /ic_def.
  iSpecialize ("He1" $! _ _ _ _ with "[$Hσ3]"); auto.
  iCombine "HQ'" "He1" as "He1".
  iMod "He1" as "[HQ' [He1 Hσ4]]".
  iDestruct "HQ'" as %HQ'.
  apply rtc_nsteps in HQ'. destruct HQ' as [n1' HQ'].
  iDestruct "He1" as (σ4'' w1'') "(Hrd4 & Hσ4' & #Hww1)".
  iDestruct "Hrd4" as %Hrd4.
  apply rtc_nsteps in Hrd4. destruct Hrd4 as [n4 Hrd4].
  destruct (nsteps_deterministic' Hrd4 HQ') as [? [? ?]]; subst.
  iMod (heap_catch_up with "Hσ4'") as "Hσ5'"; eauto.
  iDestruct "Hww1" as "[Hww1|Hww1]".
  - iDestruct "Hww1" as ([u1 u1']) "[Heq Hww1]".
    iDestruct "Heq" as %Heq; inversion Heq; subst.
    apply case_l_red_step in Hrd22.
    destruct Hrd22 as (m' & Hm' & Hrd22). asimpl in Hrd22.
    erewrite typed_subst_head_simpl_2 in Hrd22; eauto;
            last by rewrite /= !fmap_length Hvvsl.
    iPoseProof (binary_fundamental _ _ _ Ht1 Δ
      ((u1, u1') :: (w2, w2'') :: vvs) with "[]")
      as "He2"; first by rewrite !interp_env_cons; iFrame "#".
    iSpecialize ("He2" $! _ _ _ with "[$Hσ5']"). asimpl.
    rewrite ic_eq /ic_def.
    iSpecialize ("He2" $! _ _ _ _ with "[$Hσ4]"); auto.
    iMod "He2" as "[He2 Hσ5]".
    iDestruct "He2" as (σ6' w3') "(Hrd5 & Hσ6' & #Hww3)".
    iDestruct "Hrd5" as %Hrd5.
    apply rtc_nsteps in Hrd5. destruct Hrd5 as [n5 Hrd5].
    iModIntro; iFrame.
    iExists _, _; iFrame; iFrame "#".
    iPureIntro.
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
      eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
      eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
    erewrite typed_subst_head_simpl; eauto;
      last by rewrite /= !fmap_length Hvvsl.
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (rtc_l _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (rtc_l _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      rewrite lookup_swap_201. asimpl. econstructor. }
    erewrite typed_subst_head_simpl_2; eauto;
      last by rewrite /= !fmap_length Hvvsl.
    eapply (nsteps_rtc _); eauto.
  - iDestruct "Hww1" as ([u1 u1']) "[Heq Hww1]".
    iDestruct "Heq" as %Heq; inversion Heq; subst.
    apply case_r_red_step in Hrd22.
    destruct Hrd22 as (m' & Hm' & Hrd22). asimpl in Hrd22.
    erewrite typed_subst_head_simpl_2 in Hrd22; eauto;
            last by rewrite /= !fmap_length Hvvsl.
    iPoseProof (binary_fundamental _ _ _ Ht2 Δ
      ((u1, u1') :: (w2, w2'') :: vvs) with "[]")
      as "He2"; first by rewrite !interp_env_cons; iFrame "#".
    iSpecialize ("He2" $! _ _ _ with "[$Hσ5']"). asimpl.
    rewrite ic_eq /ic_def.
    iSpecialize ("He2" $! _ _ _ _ with "[$Hσ4]"); auto.
    iMod "He2" as "[He2 Hσ5]".
    iDestruct "He2" as (σ6' w3') "(Hrd5 & Hσ6' & #Hww3)".
    iDestruct "Hrd5" as %Hrd5.
    apply rtc_nsteps in Hrd5. destruct Hrd5 as [n5 Hrd5].
    iModIntro; iFrame.
    iExists _, _; iFrame; iFrame "#".
    iPureIntro.
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
      eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
      eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
    erewrite typed_subst_head_simpl; eauto;
      last by rewrite /= !fmap_length Hvvsl.
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (CaseCtx _ _)));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (rtc_l _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      asimpl. econstructor. }
    etrans.
    { eapply (nsteps_rtc _).
      eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
        eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
    etrans.
    { eapply (rtc_l _ (_, _) (_, _)); simpl.
      { apply head_prim_step. econstructor; eauto. }
      rewrite lookup_swap_201. asimpl. econstructor. }
    erewrite typed_subst_head_simpl_2; eauto;
      last by rewrite /= !fmap_length Hvvsl.
    eapply (nsteps_rtc _); eauto.
Qed.