RunST.tlam_beta

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.

Lemma tlam_beta_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
  Γ e log e' : TForall τ
  Γ TApp (TLam e) log e' : TForall τ.
Proof.
  iIntros (HLR Δ vvs ) "#(Hinv & Hvvs)".
  iDestruct (interp_env_length with "[]") as %Hvvsl; iFrame "#".
  iIntros (γh γh' σ1') "Hσ1' /=".
  rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
  iDestruct "Hrd1" as %Hrd1.
  apply TLam_red_step in Hrd1. destruct Hrd1 as (m & Hm & Hrd1); subst.
  iPoseProof (HLR with "[]") as "He"; iFrame "#".
  iSpecialize ("He" $! _ _ _ with "[$Hσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He" $! _ _ _ _ with "[$How]"); eauto.
  iMod "He" as "[He HPh1]".
  iDestruct "He" as (σ2'' v'' ) "(Hrd2 & Ho1 & #Hrl1)".
  iDestruct "Hrd2" as %Hrd2.
  iModIntro; iFrame.
  iExists _, _; iFrame; iSplit; eauto.
Qed.

Lemma tlam_beta_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
  Γ e log e' : TForall τ
  Γ e log TApp (TLam e') : TForall τ.
Proof.
  iIntros (HLR Δ 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.
  iPoseProof (HLR with "[]") as "He"; iFrame "#".
  iSpecialize ("He" $! _ _ _ with "[$Hσ1']"); simpl.
  rewrite ic_eq /ic_def /=.
  iSpecialize ("He" $! _ _ _ _ with "[$How]"); eauto.
  iMod "He" as "[He HPh1]".
  iDestruct "He" as (σ2'' v'' ) "(Hrd2 & Ho1 & #Hrl1)".
  iDestruct "Hrd2" as %Hrd2.
  iModIntro; iFrame.
  apply rtc_nsteps in Hrd2. destruct Hrd2 as [k Hrd2].
  iExists σ2'', _; iSplit; iFrame.
  { iPureIntro.
    eapply (rtc_l _ (_, _) (_, _)); simpl.
    { apply head_prim_step. econstructor; eauto. }
    eapply (nsteps_rtc _); eauto. }
  auto.
Qed.

Lemma tlam_beta Γ e τ :
  Γ ⊢ₜ e : TForall τ
  Γ TApp (TLam e) ctx e : TForall τ.
Proof.
  intros Ht. split.
  - apply binary_soundness; auto.
    + replace (TForall τ) with (TForall τ).[ren (+1)].[TUnit/] by by asimpl.
      econstructor. econstructor. by apply context_rename.
    + intros. eapply tlam_beta_LR1; eauto using binary_fundamental.
  - apply binary_soundness; auto.
    + replace (TForall τ) with (TForall τ).[ren (+1)].[TUnit/] by by asimpl.
      econstructor. econstructor. by apply context_rename.
    + intros. eapply tlam_beta_LR2; eauto using binary_fundamental.
Qed.