RunST.ST_Lang_reduction

From iris.program_logic Require Import ectx_language ectx_lifting.
From iris.program_logic Require ectxi_language.
From RunST Require Import base lang reduction.

 Lemma to_val_of_eff_val ev v :
  to_val (of_eff_val ev) = Some v w, ev = ReturnEV w.
 Proof.
  destruct ev as [|K]; simpl; eauto.
  intros HK; exfalso; revert v HK; induction K as [|[]] => w HK; first done.
  simpl in *.
  destruct (to_val (foldr eff_fill_item (App (of_val v1) (of_val v2)) K));
   eauto.
 Qed.

Lemma eff_red_val e σ :
  @nf_head_reducible _ _ _ _ eff_ectx_lang e σ is_Some (to_val e).
Proof.
  destruct e; inversion 1 as (e' & σ' & H'); inversion H'; subst; simpl;
    repeat match goal with H : to_val _ = Some _ |- _ => rewrite H end; simpl;
    eauto.
Qed.

Lemma eff_head_step_kind e σ e' σ' :
  @eff_head_step e σ e' σ' []
  ( v1 v2, e' = App (of_val v1) (of_val v2)) v, e' = Return (of_val v).
Proof.
  inversion 1; subst; eauto.
  right; eexists (LitV _); eauto.
  right; eexists (LitV _); eauto.
  left; do 2 eexists.
  repeat match goal with H : to_val _ = _ |- _ => rewrite -(of_to_val _ _ H) end.
  eauto.
Qed.

Lemma fill_eff_res_kind (K : list eff_ectx_item) e :
 (( v1 v2, e = App (of_val v1) (of_val v2)) v, e = Return (of_val v))
 is_Some (to_eff_val (fill K e)) is_Some (to_val (fill K e)).
Proof.
  intros [(e1 & e2 & He)|[e' He]]; subst.
  - left. induction K as [|[] K]; simpl; rewrite !to_of_val; simpl; eauto.
    destruct IHK as [[] HK]; rewrite HK; simpl.
    destruct K as [|[] K]; simpl in *; inversion HK.
    rewrite !to_of_val in HK; inversion HK.
    destruct (to_eff_val (fill K (App (of_val e1) (of_val e2))));
      inversion HK; simpl in *.
    rewrite to_of_val in HK; simpl in *.
    destruct e; simpl in *; inversion HK.
    eauto.
  - right. induction K as [|[] K]; simpl; rewrite to_of_val; simpl; eauto.
    destruct IHK as [? HK]; rewrite HK; eauto.
Qed.

Lemma eff_prim_step_either_val e σ e' σ' :
  @prim_step eff_lang e σ e' σ' []
  is_Some (to_eff_val e') is_Some (to_val e').
Proof.
  intros Hpr; inversion Hpr as [K ? ? ? ? Hhr]; subst.
  apply eff_head_step_kind in Hhr; by apply fill_eff_res_kind.
Qed.

Lemma red_runST n v v' σ σ' :
  nsteps pstep n (RunST (of_val v), σ) (of_val v', σ')
   k w σ'',
    k n
    nsteps (pstep' (eff_lang)) k ((of_val v), σ) (of_eff_val w, σ'')
    nsteps pstep (n-k) (RunST (of_eff_val w), σ'') (of_val v', σ').
Proof.
  revert v σ; induction n => v σ Hr.
  - inversion Hr as [Hs1 Hs2 Hs3 Hs4|]; subst.
    assert (Hkv : to_val (RunST (of_val v)) = Some v') by
        by rewrite Hs4 to_of_val.
    exfalso; by induction v.
  - inversion Hr as [|Hs1 Hs2 [e2 σ2] Hs4 Hs5 Hs6]; subst.
    simpl in *.
    inversion Hs5 as [? ? ? Hfl ? Hhd]; subst.
    destruct K as [|[] K]; simpl in *; inversion Hfl; subst.
    + inversion Hhd; subst.
      * destruct (eff_prim_step_either_val _ _ _ _ H1) as [[w Hw]|[w Hw]].
        -- exists 1, w, σ2. repeat split; first omega.
           ++ eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
              rewrite -(of_to_eff_val _ _ Hw); econstructor.
           ++ simpl; replace (n - 0) with n by omega.
              by rewrite (of_to_eff_val _ _ Hw).
        -- rewrite -(of_to_val _ _ Hw) in Hs6.
           destruct (IHn w σ2 Hs6) as (k & w' & σ'' & Hle & Hstp1 & Hstp2).
           exists (S k), w', σ''; repeat split; first omega; auto.
           eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
             by rewrite -(of_to_val _ _ Hw).
      * eexists 0, (ReturnEV _), _. repeat split; first omega.
        -- simpl. erewrite of_to_val; first econstructor; eauto.
        -- replace (S n - 0) with (S n) by omega. simpl.
           by rewrite (of_to_val _ _ H1); rewrite H0.
    + apply val_stuck in Hhd.
      pose proof (to_of_val v) as Hnv; rewrite H0 in Hnv.
      by erewrite (@fill_not_val _ _ _ _ ST_ectxi_lang) in Hnv.
Qed.

Lemma red_RunST_Bind n e w u σ σ' :
  nsteps pstep n (RunST (Bind e (of_val w)), σ) (of_val u, σ')
   k u' σ'',
    k n
    nsteps pstep k (RunST e, σ) (of_val u', σ'')
    nsteps pstep (n-k) (RunST (App (of_val w) (of_val u')), σ'') (of_val u, σ').
Proof.
  revert e σ; induction n => e σ Hr.
  - inversion Hr as [Hs1 Hs2 Hs3 Hs4|]; subst.
    assert (Hkv : to_val (RunST (Bind e (of_val w))) = Some u) by
        by rewrite Hs4 to_of_val.
    exfalso; by induction e.
  - inversion Hr as [|Hs1 Hs2 [e2 σ2] Hs4 Hs5 Hs6]; subst.
    simpl in *.
    inversion Hs5 as [? ? ? Hfl ? Hhd]; subst.
    destruct K as [|[] K]; simpl in *; inversion Hfl; subst.
    + inversion Hhd; subst.
      match goal with
        H : prim_step (Bind e (of_val w)) _ _ _ _ |- _ => inversion H; subst
      end.
      destruct K as [|[] K]; simpl in *; inversion H0; subst.
      * inversion H3; subst.
        eexists 1, _, _; repeat split; first omega.
        { eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
          - eapply head_prim_step; eapply RunRet; eauto.
          - rewrite -(of_to_val _ _ H5); econstructor. }
        simpl. replace (n - 0) with n by omega. by rewrite (of_to_val _ _ H5).
      * rewrite -(of_val_inj _ _ H5) in Hs6.
        destruct (IHn _ _ Hs6) as (k & u' & σ'' & Hle & Hrd1 & Hrd2).
        eexists (S k), _, _; repeat split; eauto with omega.
        econstructor; eauto; simpl.
        apply head_prim_step; econstructor. econstructor; eauto.
    + destruct K as [|[] K]; simpl in *; inversion H0; subst.
      * inversion Hhd.
      * destruct (IHn _ _ Hs6) as (k & u' & σ'' & Hle & Hrd1 & Hrd2).
        eexists (S k), _, _; repeat split; eauto with omega.
        econstructor; eauto; simpl.
        eapply (Ectx_step _ _ _ _ _ (RunSTCtx :: _)); simpl; eauto.
      * apply val_stuck in Hhd.
        pose proof (to_of_val w) as Hnv; rewrite H2 in Hnv.
        by erewrite (@fill_not_val _ _ _ _ ST_ectxi_lang) in Hnv.
Qed.

Lemma red_RunST_Bind' n e w u σ σ' k u' σ'' :
  nsteps (pstep' det_lang) n (RunST e, σ) (of_val u', σ'')
  nsteps (pstep' det_lang) k (RunST (App (of_val w) (of_val u')), σ'') (of_val u, σ')
  nsteps (pstep' det_lang) (n + k) (RunST (Bind e (of_val w)), σ) (of_val u, σ').
Proof.
  revert k w e u u' σ σ' σ'';
    induction n => k w e u u' σ σ' σ'' Hrd1 Hrd2 /=.
  - inversion Hrd1 as [Hs1 Hs2 Hs3 Hs4|]; subst.
    assert (Hkv : to_val (RunST e) = Some u') by
        by rewrite Hs4 to_of_val.
    exfalso; by induction e.
  - inversion Hrd1 as [|Hs1 Hs2 [e2 σ2] Hs4 Hs5 Hs6]; subst; simpl in *.
    inversion Hs5 as [? ? ? Hfl ? Hhd]; subst.
    destruct K as [|[] K]; simpl in *; inversion Hfl; subst.
    + inversion Hhd; subst.
      * specialize (IHn _ _ _ _ _ _ _ _ Hs6 Hrd2).
        eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
        apply head_prim_step.
        econstructor.
        inversion H1 as [K ? ? ?]; subst.
        eapply (Ectx_step _ _ _ _ _ (BindECtx _ :: K)); simpl; eauto.
      * inversion Hs6; subst; simpl in *.
        -- eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
           apply head_prim_step.
           econstructor.
           apply head_prim_step.
           econstructor; eauto using to_of_val.
        -- apply val_prim_stuck in H0; simpl in *; by rewrite H1 in H0.
    + specialize (IHn _ _ _ _ _ _ _ _ Hs6 Hrd2).
      eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
      eapply (Ectx_step _ _ _ _ _ (RunSTCtx :: BindLCtx _ :: _)); simpl; eauto.
Qed.

Lemma nsteps_det n e σ e' σ' :
  nsteps (pstep' det_lang) n (e, σ) (e', σ')
  nsteps (pstep' ST_lang) n (e, σ) (e', σ').
Proof. eapply nsteps_ind; econstructor; eauto using prim_step_impl. Qed.

Lemma prim_step_deterministic {e σ e1' σ1' e2' σ2'} :
  @prim_step det_lang e σ e1' σ1' []
  @prim_step det_lang e σ e2' σ2' []
  e1' = e2' σ1' = σ2'.
Proof.
  intros Hrd1 Hrd2.
  inversion Hrd1 as [K e3 e4 ? ? Hhrd1]; subst.
  inversion Hrd2 as [K' e5 e6 HKeq HKeq' Hhrd2]; subst.
  replace K' with K in *; last first.
  { eapply (@step_by_val_strong _ _ _ _ det_ectxi_lang); simpl in *;
      first apply HKeq; eauto. }
  replace e3 with e5 in * by by
        rewrite (@fill_inj _ _ _ _ det_ectxi_lang _ _ _ HKeq).
  edestruct det_head_step_det as [Heq1 [He2 ?]]; first apply Hhrd1; first apply Hhrd2.
  by subst.
Qed.

Lemma nsteps_deterministic {n e σ e1' σ1' e2' σ2'} :
  nsteps (pstep' det_lang) n (e, σ) (e1', σ1')
  nsteps (pstep' det_lang) n (e, σ) (e2', σ2')
  e1' = e2' σ1' = σ2'.
Proof.
  revert e σ e1' σ1' e2' σ2'; induction n => e σ e1' σ1' e2' σ2' Hrd1 Hrd2.
  - by inversion Hrd1; inversion Hrd2; subst.
  - inversion Hrd1 as [|Hs11 Hs12 [e12 σ12] Hs14 Hs15 Hs16]; subst; simpl in *.
    inversion Hrd2 as [|Hs21 Hs22 [e22 σ22] Hs24 Hs25 Hs26]; subst; simpl in *.
    destruct (prim_step_deterministic Hs15 Hs25); subst.
    eapply IHn; eauto.
Qed.

Lemma nsteps_deterministic' {n n' e σ v1 σ1' v2 σ2'} :
  nsteps (pstep' det_lang) n (e, σ) (of_val v1, σ1')
  nsteps (pstep' det_lang) n' (e, σ) (of_val v2, σ2')
  n = n' v1 = v2 σ1' = σ2'.
Proof.
  revert n' e σ v1 σ1' v2 σ2'; induction n => n' e σ v1 σ1' v2 σ2' Hrd1 Hrd2.
  - inversion Hrd1 as [Hs11 Hs12 Hs13 Hs14|Hs11 Hs12 [e12 σ12] Hs14 Hs15 Hs16];
      inversion Hrd2 as [Hs21 Hs22 Hs23 Hs24|Hs21 Hs22 [e22 σ22] Hs24 Hs25 Hs26];
      subst; simpl in *.
    + by apply of_val_inj in Hs24; subst.
    + by apply val_prim_stuck in Hs25; simpl in Hs25; rewrite to_of_val in Hs25.
  - inversion Hrd1 as [Hs11 Hs12 Hs13 Hs14|Hs11 Hs12 [e12 σ12] Hs14 Hs15 Hs16];
      inversion Hrd2 as [Hs21 Hs22 Hs23 Hs24|k Hs22 [e22 σ22] Hs24 Hs25 Hs26];
      subst; simpl in *.
    + by apply val_prim_stuck in Hs15; simpl in Hs15; rewrite to_of_val in Hs15.
    + destruct (prim_step_deterministic Hs15 Hs25); subst.
      cut (n = k v1 = v2 σ1' = σ2'); first intuition auto.
      eapply IHn; eauto.
Qed.

Hint Resolve to_of_val.

Lemma head_red_app k e w v σ1 σ2:
  nsteps (pstep' ST_lang) k (App (Rec (Rec e)) (of_val w), σ1) (of_val v, σ2)
  k = 1 v = (RecV e.[upn 2 (Rec (Rec e) .: of_val w .: ids)])
   σ1 = σ2.
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct v; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *.
    inversion Hs16 as [Hs31 Hs32 Hs33 Hs34|k'' Hs42 [e42 σ42] Hs44 Hs45 Hs46];
      subst; simpl in *.
    + repeat split; trivial. apply of_val_inj. rewrite -Hs34.
      by asimpl.
    + apply val_prim_stuck in Hs45. inversion Hs45.
Qed.

Lemma rec_red_step n e v σ w σ':
  nsteps (pstep' ST_lang) n (App (Rec e) (of_val v), σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e.[(Rec e), of_val v/], σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma det_rec_red_step n e v σ w σ':
  nsteps (pstep' det_lang) n (App (Rec e) (of_val v), σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' det_lang) m (e.[(Rec e), of_val v/], σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma head_red_app_TLam k e w v σ1 σ2:
  nsteps (pstep' ST_lang) k (App (Rec (TLam e)) (of_val w), σ1) (of_val v, σ2)
  k = 1 v = (TLamV e.[(Rec (TLam e) .: of_val w .: ids)])
   σ1 = σ2.
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct v; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *.
    inversion Hs16 as [Hs31 Hs32 Hs33 Hs34|k'' Hs42 [e42 σ42] Hs44 Hs45 Hs46];
      subst; simpl in *.
    + repeat split; trivial. apply of_val_inj. rewrite -Hs34.
      by asimpl.
    + apply val_prim_stuck in Hs45. inversion Hs45.
Qed.

Lemma TLam_red_step n e σ w σ':
  nsteps (pstep' ST_lang) n (TApp (TLam e), σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e, σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma case_l_red_step n v e1 e2 σ w σ':
  nsteps (pstep' ST_lang) n (Case (InjL (of_val v)) e1 e2, σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e1.[of_val v/], σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma case_r_red_step n v e1 e2 σ w σ':
  nsteps (pstep' ST_lang) n (Case (InjR (of_val v)) e1 e2, σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e2.[of_val v/], σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma if_true_red_step n e1 e2 σ w σ':
  nsteps (pstep' ST_lang) n (If (Lit $ LitBool true) e1 e2, σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e1, σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.

Lemma if_false_red_step n e1 e2 σ w σ':
  nsteps (pstep' ST_lang) n (If (Lit $ LitBool false) e1 e2, σ) (of_val w, σ')
   m, n = S m
       nsteps (pstep' ST_lang) m (e2, σ) ((of_val w), σ').
Proof.
  intros Hrd.
  inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
    subst; simpl in *.
  - destruct w; inversion Hs14.
  - eapply head_reducible_prim_step in Hs15; last repeat econstructor; eauto.
    inversion Hs15; subst. asimpl in *. eauto.
Qed.