RunST.rec_hoisting.rec_hoisting_part3

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

Hint Resolve to_of_val.

Require Import Coq.Logic.FunctionalExtensionality.

Lemma lookup_swap_3 : (lookup_n (swap_list 1 2)) = (1 .: 2 .: 0 .: (+3)).
Proof.
  extensionality i; destruct i as [|[|[|?]]]; by cbv.
Qed.

Lemma Rec_Hoisting_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' τ' τ'' :
  Γ ⊢ₜ e1' : τ [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ''
  Γ App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
                    e1'.[ren (+2)])).[ren (+2)]) e1' log
      Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
           e1'.[ren (+2)]) : TArrow τ' τ''.
Proof.
  rewrite lookup_swap_3. asimpl.
  iIntros (Ht1 Ht2 Δ vvs ) "#(Hinv & Hvvs)".
  iIntros (γh γh' σ1') "Hσ1'". asimpl.
  rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
  iDestruct "Hrd1" as %Hrd1.
  apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
  destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12). simpl in *.
  iPoseProof (binary_fundamental _ _ _ Ht1 with "[]") as "He1";
    iFrame "#".
  iMod (allocate_full_state σ1') as (γh3) "[Hnfσ1' Hnoσ1']".
  iSpecialize ("He1" $! _ _ _ with "[$Hnfσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
  iMod "He1" as "[Hrd How]". iDestruct "Hrd" as (σ2' w') "(Hrd2 & Hσ2' & #Hww)".
  iDestruct "Hrd2" as %Hrd2.
  apply rtc_nsteps in Hrd2. destruct Hrd2 as [k' Hrd2].
  apply head_red_app in Hrd12; destruct Hrd12 as [? [? ?]]; subst.
  iModIntro; iFrame. iExists _, (RecV _); iFrame; iSplit.
  { iPureIntro. asimpl. econstructor 1. }
  asimpl.
  replace
    (Rec (App (Rec
       e2'.[ids 2 .: ids 3 .: ids 1 .: env_subst (vvs.*1) >> ren (+4)])
         e1'.[env_subst (vvs.*1) >> ren (+2)]))
  with
  (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
            e1'.[ren (+2)])).[env_subst (vvs.*1)]
    by by rewrite lookup_swap_3; asimpl.
  replace
    (Rec (App (Rec
       e2'.[ids 2 .: ids 3 .: ids 1 .: env_subst (vvs.*2) >> ren (+4)])
         e1'.[env_subst (vvs.*2) >> ren (+2)]))
  with
  (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
            e1'.[ren (+2)])).[env_subst (vvs.*2)]
    by by rewrite lookup_swap_3; asimpl.
  assert (Γ ⊢ₜ (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
            e1'.[ren (+2)])) : TArrow τ' τ'') as Ht2'.
  { constructor. econstructor; last by eapply (context_weakening [_; _]).
    econstructor. eapply (context_weakening [_]).
    eapply (context_swap_typed [] [_; _] [_]); auto. }
  iClear "Hnoσ1' Hσ2'". iAlways.
  iIntros (vv) "#Hvv".
  iPoseProof (binary_fundamental _ _ _ Ht2' with "[]") as "Hff";
    first by iFrame "#".
  iApply (logrel_bind [AppLCtx _] [AppLCtx _]).
  iSplitL; first iExact "Hff".
  by iIntros (ff) "Hff"; iApply "Hff".
Qed.