RunST.lam_beta.lam_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 rec_hoisting.rec_hoisting
     rec_hoisting.rec_hoisting_part2
     tlam_hoisting.

From RunST.lam_beta Require Import lam_beta_part1 lam_beta_part2
     lam_beta_part3 lam_beta_part4.

From iris.proofmode Require Import tactics.

Hint Resolve to_of_val.

Lemma subst_ren_0 (Γ : list type) :
subst (ren (+0)) <$> Γ = Γ.
Proof.
  induction Γ; asimpl in *; auto using f_equal.
Qed.

Lemma subst_ren_Sn n (Γ : list type) :
subst (ren (+S n)) <$> Γ = subst (ren S) <$> subst (ren (+n)) <$> Γ.
Proof.
  induction Γ; asimpl; auto.
  by f_equal.
Qed.

Lemma subst_ren_S_inj (Γ Γ' : list type) :
  subst (ren (+1)) <$> Γ = subst (ren (+1)) <$> Γ' Γ = Γ'.
Proof.
  revert Γ'; induction Γ => Γ'; destruct Γ'; inversion 1; asimpl; auto.
  apply lift_inj in H1; subst.
  f_equal; auto.
Qed.

Lemma subst_ren_n_inj n (Γ Γ' : list type) :
  subst (ren (+n)) <$> Γ = subst (ren (+n)) <$> Γ' Γ = Γ'.
Proof.
  induction n.
  - asimpl. apply inj; typeclasses eauto.
  - rewrite !subst_ren_Sn; intros Hs%subst_ren_S_inj; auto.
Qed.

Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_n_0 : lookup_n [0] = id.
Proof.
  extensionality i; destruct i; auto.
Qed.

Lemma lookup_simpl_for_lam_beta_rec (ξ : list type) :
  (2 .: 3 .: lookup_n (swap_list 1 (length ξ)) >>> 1 .: (+4)) =
  (lookup_n (1 :: 2 :: seq 3 (length ξ) ++ [0]) >>> (+1)) :> (var var).
Proof.
  extensionality i. asimpl.
  destruct i as [|[| i]]; auto. asimpl.
  unfold lookup_n, swap_list. asimpl.
  destruct (decide (i < (length ξ))) as [Hi|Hi].
  - repeat (rewrite lookup_app_l; last by rewrite seq_length).
    rewrite !lookup_seq; auto.
  - repeat (rewrite lookup_app_r; last by rewrite seq_length; auto with omega).
    rewrite !seq_length.
    destruct (i - (length ξ)) eqn:Heq; auto with omega.
    asimpl. destruct i; auto with omega.
Qed.

Lemma subst_simpl_for_lam_beta_case (e : expr) n :
  e.[ren (lookup_n (swap_list 1 (S n)))].[ren (+1)] =
  e.[up (ren (lookup_n (swap_list 1 n)))].[up (ren (+1))]
   .[ren (lookup_n (swap_list 2 1))].
Proof.
  asimpl.
  replace (@ren expr _ (lookup_n (1 :: seq 2 n ++ [0]) >>> (+1))) with
  (@ren expr _ (2 .: lookup_n (swap_list 1 n) >>> (+2) >>> lookup_n [2; 0; 1]));
    first done.
  extensionality i.
  asimpl. rewrite /lookup_n /swap_list.
  destruct i; asimpl; first done.
  destruct (decide (i < n)).
  - rewrite !lookup_app_l ?seq_length //= !lookup_seq //=.
  - rewrite !lookup_app_r ?seq_length; auto with omega.
    destruct (i - n) as [|k] eqn:Heq; asimpl; auto.
    destruct i; auto with omega.
Qed.

Lemma lam_beta_base n ξ ctx Γ e1' e2' τ τ' :
  ctx ⊢ₜ e1' : τ'
  ctx = ξ ++ subst (ren (+n)) <$> τ :: Γ
   ξ ++ subst (ren (+n)) <$> Γ ⊢ₜ e2'.[ren (+length ξ)] : τ.[ren (+n)]
     ξ ++ subst (ren (+n)) <$> Γ
        App (Rec e1'.[ren (lookup_n (swap_list 1 (length ξ)))].[ren (+1)])
            e2'.[ren (+length ξ)]
      ctx e1'.[upn (length ξ) (e2' .: ids)] : τ'.
Proof.
  intros Ht. revert n ξ.
  induction Ht => n ξ Hctxeq Ht'; subst.
  - eapply lam_beta_var; eauto.
  - asimpl. eapply (lam_beta_basic_type (LitV _)); simpl; eauto using typed.
    iIntros (? ? ? ?) ""; by iSplit.
  - asimpl. eapply (lam_beta_basic_type (LitV _)); simpl; eauto using typed.
    iIntros (? ? ? ?) "". by iExists _; iSplit.
  - asimpl. eapply (lam_beta_basic_type (LitV _)); simpl; eauto using typed.
    iIntros (? ? ? ?) "". by iExists _; iSplit.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BinOpL _ _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BinOpR _ _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 (BinOp op)); eauto using typed.
    + intros; eexists (BinOpLCtx _ _); eauto.
    + intros; eexists (BinOpRCtx _ _); eauto.
    + intros; eapply val_rel_expr_rel_binop; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_PairL _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_PairR _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 Pair); eauto using typed.
    + intros; eexists (PairLCtx _); eauto.
    + intros; eexists (PairRCtx _); eauto.
    + intros; eapply val_rel_expr_rel_pair; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Fst]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam FstCtx); eauto using typed.
      intros; eapply bin_log_related_fst; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Snd]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam SndCtx); eauto using typed.
      intros; eapply bin_log_related_snd; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_InjL]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam InjLCtx); eauto using typed.
      intros; eapply bin_log_related_injl; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_InjR]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam InjRCtx); eauto using typed.
      intros; eapply bin_log_related_injr; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_CaseL _ _]); eauto.
      repeat econstructor; eauto; asimpl.
      - eapply (typed_gen_subst (_ :: _)); eauto.
        eapply context_strengthening; eauto.
      - eapply (typed_gen_subst (_ :: _)); eauto.
        eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_CaseM _ _]); eauto;
      last apply (IHHt2 n (τ1 :: ξ) eq_refl).
      - repeat econstructor; eauto.
        + eapply (context_weakening [_]).
          eapply (context_swap_typed [] _ [_]); eauto.
        + eapply (typed_gen_subst (_ :: _)); eauto.
          eapply context_strengthening; eauto.
      - replace e2'.[ren (+length (τ1 :: ξ))] with
        e2'.[ren (+length (ξ))].[ren (+1)] by by asimpl.
        eapply (context_weakening [_]); eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_CaseR _ _]); eauto;
      last apply (IHHt3 n (τ2 :: ξ) eq_refl).
      - repeat econstructor; eauto.
        + eapply (context_weakening [_]).
          eapply (context_swap_typed [] _ [_]); eauto.
        + eapply (context_weakening [_]).
          eapply (context_swap_typed [] (_ :: _) [_]); eauto.
        + replace e2'.[ren (+length (τ1 :: ξ))] with
        e2'.[ren (+length (ξ))].[ren (+1)] by by asimpl.
        eapply (context_weakening [_]); eauto.
      - replace e2'.[ren (+length (τ1 :: ξ))] with
        e2'.[ren (+length (ξ))].[ren (+1)] by by asimpl.
        eapply (context_weakening [_]); eauto. }
    simpl. rewrite !subst_simpl_for_lam_beta_case.
    replace e2'.[ren (+S (length ξ))] with e2'.[ren (+(length ξ))].[ren (+1)]
    by by asimpl.
    eapply lam_beta_push_lam_case; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [_] _ [_]); eauto.
    + eapply (context_swap_typed [_] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_IfL _ _]); eauto.
      repeat econstructor; eauto; asimpl.
      - eauto using typed_gen_subst, context_strengthening.
      - eauto using typed_gen_subst, context_strengthening. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_IfM _ _]); eauto.
      repeat econstructor; eauto.
      + eapply (context_weakening [_]).
        eapply (context_swap_typed [] _ [_]); eauto.
      + eauto using typed_gen_subst, context_strengthening. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_IfR _ _]); eauto.
      repeat econstructor; eauto.
      + eapply (context_weakening [_]).
         eapply (context_swap_typed [] _ [_]); eauto.
      + eapply (context_weakening [_]).
        eapply (context_swap_typed [] _ [_]); eauto. }
    eapply lam_beta_push_lam_if; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Rec]); eauto.
      repeat econstructor; eauto.
      specialize (IHHt n (TArrow τ1 τ2 :: τ1 :: ξ) eq_refl).
      asimpl in *.
      apply IHHt; eauto.
      replace (e2'.[ren (+S (S (length ξ)))]) with
      e2'.[ren (+(length ξ))].[ren (+2)] by by asimpl.
      by apply (context_weakening [_;_]).
    + replace (Rec e).[ren (lookup_n (swap_list 1 (length ξ)))].[ren (+1)] with
      (Rec e.[upn 2 (ren (lookup_n (swap_list 1 (length ξ))))]).[ren (+1)]
        by by asimpl.
      replace (e2'.[ren (+S (S (length ξ)))]) with
        e2'.[ren (+(length ξ))].[ren (+2)] by by asimpl.
      replace (e.[ren (lookup_n (1 :: 2 :: seq 3 (length ξ) ++ [0]) >>> (+1))])
        with
        e.[upn 2 (ren (lookup_n (swap_list 1 (length ξ))))]
         .[ren (lookup_n (swap_list 1 2))].[ren (+1)]; last first.
      { rewrite lookup_swap_3. asimpl.
        by rewrite lookup_simpl_for_lam_beta_rec. }
      eapply Rec_Hoisting; eauto.
      apply (context_swap_typed [_;_] _ [_]); auto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_AppL _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_AppR _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 App); eauto using typed.
    + intros; eexists (AppLCtx _); eauto.
    + intros; eexists (AppRCtx _); eauto.
    + intros; eapply val_rel_expr_rel_app; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_TLam]); eauto.
      repeat econstructor; eauto.
      specialize (IHHt (S n) (subst (ren (+1)) <$> ξ)).
      rewrite fmap_app in IHHt.
      rewrite -subst_ren_Sn in IHHt.
      specialize (IHHt eq_refl).
      rewrite !fmap_length in IHHt.
      rewrite fmap_app -subst_ren_Sn.
      apply IHHt; eauto.
      rewrite (subst_ren_Sn n) -fmap_app.
      replace τ.[ren (+S n)] with τ.[ren (+n)].[ren (+1)] by by asimpl.
      apply (context_rename); auto.
    + replace (TLam e).[ren (lookup_n (swap_list 1 (length ξ)))].[ren (+1)] with
      (TLam e.[ren (lookup_n (swap_list 1 (length ξ)))]).[ren (+1)]
        by by asimpl.
      eapply TLam_Hoisting; eauto.
      cbn -[subst]; rewrite fmap_app.
      rewrite fmap_app in Ht.
      erewrite <- fmap_length.
      apply (context_swap_typed [] _ [_]); auto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_TApp]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam TAppCtx); eauto using typed.
      intros; eapply bin_log_related_tapp; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Fold]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam FoldCtx); eauto using typed.
      intros; eapply bin_log_related_fold; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Unfold]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam UnfoldCtx); eauto using typed.
      intros; eapply bin_log_related_unfold; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Alloc]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam AllocCtx); eauto using typed.
      intros; eapply bin_log_related_alloc; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Read]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam ReadCtx); eauto using typed.
      intros; eapply bin_log_related_read; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_WriteL _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_WriteR _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 Write); eauto using typed.
    + intros; eexists (WriteLCtx _); eauto.
    + intros; eexists (WriteRCtx _); eauto.
    + intros; eapply val_rel_expr_rel_write; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindL _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindR _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 Bind); eauto.
    + intros; eexists (BindLCtx _); eauto.
    + intros; eexists (BindRCtx _); eauto.
    + intros. econstructor; eauto.
    + intros; eapply val_rel_expr_rel_bind; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Return]); eauto.
      repeat econstructor; eauto.
    + eapply (lam_beta_push_lam ReturnCtx); eauto using typed.
      intros; eapply bin_log_related_return; eauto.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    + eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_RunST]); eauto.
      repeat econstructor; eauto.
      specialize (IHHt (S n) (subst (ren (+1)) <$> ξ)).
      rewrite fmap_app in IHHt.
      rewrite -subst_ren_Sn in IHHt.
      specialize (IHHt eq_refl).
      rewrite !fmap_length in IHHt.
      rewrite fmap_app -subst_ren_Sn.
      apply IHHt; eauto.
      rewrite (subst_ren_Sn n) -fmap_app.
      replace τ.[ren (+S n)] with τ.[ren (+n)].[ren (+1)] by by asimpl.
      apply (context_rename); auto.
    + eapply lam_beta_push_lam_RunST; eauto.
      cbn -[subst]; rewrite fmap_app.
      rewrite fmap_app in Ht.
      erewrite <- fmap_length.
      eapply (context_swap_typed [] _ [_]); eauto.
  - eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_CompL _]); eauto.
      repeat econstructor; eauto.
      eapply typed_gen_subst; eauto.
      eapply context_strengthening; eauto. }
    eapply ctx_refines_trans; last first.
    { eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_CompR _]); eauto.
      repeat econstructor; eauto.
      eapply (context_weakening [_]).
      eapply (context_swap_typed [] _ [_]); eauto. }
    eapply (lam_beta_push_lam_2 Compare); eauto using typed.
    + intros; eexists (CompLCtx _); eauto.
    + intros; eexists (CompRCtx _); eauto.
    + intros; eapply val_rel_expr_rel_compare; eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
    + eapply (context_swap_typed [] _ [_]); eauto.
Qed.

Lemma lam_beta Γ e1' e2' τ τ':
  τ :: Γ ⊢ₜ e1' : τ' ->
  Γ ⊢ₜ e2' : τ ->
  Γ App (Rec e1'.[ren (+1)]) e2' ctx e1'.[e2'/] : τ'.
Proof.
  intros Ht1 Ht2.
  pose proof (lam_beta_base 0 []:: Γ) Γ e1' e2' τ τ' Ht1) as Hlb.
  rewrite !subst_ren_0 in Hlb. asimpl in Hlb. rewrite lookup_n_0 in Hlb.
  asimpl in Hlb. by apply Hlb.
Qed.