LogrelCC.examples.list_rev

From LogrelCC Require Import prelude.
From LogrelCC Require Import lang rules_unary cl_rules.
From LogrelCC.examples Require Import stack.
From iris.proofmode Require Import tactics.

Definition list_val_closed l :=
  Forall (λ v, forall f, (of_val v).[f] = (of_val v)) l.

Fixpoint embed_list (l : list val) : val :=
  match l with
  | [] => (FoldV (InjRV UnitV))
  | x :: l' => (FoldV (InjLV (PairV x (embed_list l'))))
  end.

Lemma embed_list_closed l :
  list_val_closed l f, (of_val (embed_list l)).[f] = (of_val (embed_list l)).
Proof.
  induction l => Hlc f; first done.
  simpl; repeat f_equal; inversion Hlc; auto.
Qed.

Section list_rev.
  Context `{heapG Σ}.

  Definition list_to_stack st :=
    Rec (Case (Unfold (Var 1))
              (App (Rec (App (Var 3) (Snd (Var 2))))
                   (App (stack_push st.[ren (+3)]) (Fst (Var 0))))
              (Unit)).

  Definition list_to_stack_val st :=
    RecV (Case (Unfold (Var 1))
              (App (Rec (App (Var 3) (Snd (Var 2))))
                   (App (stack_push st.[ren (+3)]) (Fst (Var 0))))
              (Unit)).

  Lemma list_to_stack_to_val st :
    to_val (list_to_stack st) = Some (list_to_stack_val st).
  Proof. done. Qed.

  Lemma list_to_stack_subst f st :
    (list_to_stack st).[f] = list_to_stack st.[upn 2 f].
  Proof.
    unfold list_to_stack; asimpl.
    rewrite stack_push_subst upn_comp.
    by asimpl.
  Qed.

  Lemma clwp_list_to_stack l st stl :
    list_val_closed l
    is_stack stl st
     CLWP App ((list_to_stack (Loc st))) (of_val (embed_list l))
    {{_, is_stack ((rev l) ++ stl) st }}.
  Proof.
    iIntros (Hlc) "Hst".
    iLöb as "IH" forall (l Hlc stl st) "Hst".
    iApply clwp_rec; eauto.
    asimpl.
    rewrite !embed_list_closed; auto.
    rewrite !stack_push_closed; auto.
    iNext.
    destruct l as [|x l]; simpl.
    - iApply (clwp_bind [CaseCtx _ _]).
      iApply clwp_unfold; eauto.
      iNext; iApply clwp_value; eauto; simpl.
      iApply clwp_case_injr; eauto.
      iNext. iApply clwp_value; eauto.
    - iApply (clwp_bind [CaseCtx _ _]).
      iApply clwp_unfold; auto.
      iNext.
      iApply clwp_value; auto; simpl.
      iApply clwp_case_injl; eauto.
      iNext. asimpl.
      inversion Hlc as [|? ? Hlcx Hlc2]; subst.
      rewrite !Hlcx !embed_list_closed; auto.
      rewrite !stack_push_closed; auto.
      rewrite <- (of_to_val _ _ (stack_push_to_val _)).
      iApply (clwp_bind [AppRCtx _; AppRCtx (RecV _)]).
      iApply clwp_fst; eauto.
      iNext.
      iApply clwp_value; eauto; simpl.
      rewrite (of_to_val _ _ (stack_push_to_val _)).
      iApply (clwp_bind [AppRCtx (RecV _)]).
      iApply clwp_wand_l; iSplitR "Hst";
        last by iApply clwp_stack_push; eauto.
      iIntros (w) "[% Hst]"; subst; simpl.
      iApply clwp_rec; eauto.
      iNext. asimpl.
      rewrite !Hlcx !embed_list_closed; auto.
      rewrite !stack_push_closed; auto.
      iApply (clwp_bind [AppRCtx (RecV _)]).
      iApply clwp_snd; eauto using to_of_val.
      iNext. iApply clwp_value; eauto using to_of_val.
      iSpecialize ("IH" $! l with "[]"); first iPureIntro; auto.
      iSpecialize ("IH" $! (x :: stl) st with "Hst").
      iApply clwp_wand_l; iSplitR "IH"; last by iApply "IH"; eauto.
      iIntros (?) "?". by rewrite -assoc; simpl.
  Qed.

  Typeclasses Opaque list_to_stack.
  Global Opaque list_to_stack.
  Typeclasses Opaque list_to_stack_val.
  Global Opaque list_to_stack_val.

  Definition stack_to_list st :=
    Rec (Case (App (stack_pop st.[ren (+2)]) Unit)
              (Fold (InjL (Pair (Var 0) (App (Var 1) Unit))))
              (Fold (InjR Unit))).

  Definition stack_to_list_val st :=
    RecV (Case (App (stack_pop st.[ren (+2)]) Unit)
              (Fold (InjL (Pair (Var 0) (App (Var 1) Unit))))
              (Fold (InjR Unit))).

  Lemma stack_to_list_to_val st :
    to_val (stack_to_list st) = Some (stack_to_list_val st).
  Proof. done. Qed.

  Lemma stack_to_list_subst f st :
    (stack_to_list st).[f] = stack_to_list st.[upn 2 f].
  Proof.
    unfold stack_to_list; asimpl.
    rewrite stack_pop_subst.
    by asimpl.
  Qed.

  Lemma clwp_stack_to_list st stl :
    is_stack stl st
     CLWP App ((stack_to_list (Loc st))) Unit
        {{v, v = embed_list stl is_stack [] st }}.
  Proof.
    iIntros "Hst".
    iLöb as "IH" forall (stl st) "Hst".
    iApply clwp_rec; eauto.
    asimpl.
    rewrite !stack_pop_closed; auto.
    iNext.
    destruct stl as [|x stl]; simpl.
    - iApply (clwp_bind [CaseCtx _ _]).
      iApply clwp_wand_l; iSplitR;
        last by iApply (clwp_stack_pop_empty with "Hst").
      iIntros (w) "[% Hst]"; subst; simpl.
      iApply clwp_case_injr; eauto.
      iNext. iApply clwp_value; eauto.
    - iApply (clwp_bind [CaseCtx _ _]).
      iApply clwp_wand_l; iSplitR;
        last by iApply (clwp_stack_pop_non_empty with "Hst").
      iIntros (w) "[% Hst]"; subst; simpl.
      iApply clwp_case_injl; eauto.
      iNext. asimpl.
      rewrite !stack_pop_closed; auto.
      iApply (clwp_bind [PairRCtx _; InjLCtx; FoldCtx]).
      iSpecialize ("IH" $! stl st with "Hst").
      iApply clwp_wand_l; iSplitR "IH"; last by iApply "IH"; eauto.
      iIntros (w) "[% Hst]"; subst.
      iApply clwp_value; eauto; rewrite /= !to_of_val //.
  Qed.

  Typeclasses Opaque stack_to_list.
  Global Opaque stack_to_list.
  Typeclasses Opaque stack_to_list_val.
  Global Opaque stack_to_list_val.

  Definition list_rev :=
    Rec (App (Rec (App (Rec (App (stack_to_list (Var 5)) Unit))
                       (App (list_to_stack (Var 3)) (Var 3))))
             (App create_stack Unit)).

  Definition list_rev_val :=
    RecV (App (Rec (App (Rec (App (stack_to_list (Var 5)) Unit))
                       (App (list_to_stack (Var 3)) (Var 3))))
              (App create_stack Unit)).

  Lemma list_rev_to_val :
    to_val (list_rev) = Some list_rev_val.
  Proof. done. Qed.

  Lemma list_rev_closed :
     f, list_rev.[f] = list_rev.
  Proof. by intros f; asimpl. Qed.

  Lemma clwp_list_rev l :
    list_val_closed l
    (CLWP App list_rev (of_val (embed_list l))
           {{v, v = embed_list (rev l)}})%I.
  Proof.
    iIntros (Hlc).
    iApply clwp_rec; eauto using to_of_val.
    iNext. asimpl.
    rewrite !embed_list_closed; auto.
    rewrite !create_stack_closed; auto.
    rewrite stack_to_list_subst list_to_stack_subst.
    asimpl.
    iApply (clwp_bind [AppRCtx (RecV _)]).
    iApply clwp_wand_l; iSplitR; last by iApply clwp_create_stack; eauto.
    iIntros (w) "H". iDestruct "H" as (st) "[% Hst]"; subst.
    iApply clwp_rec; eauto using to_of_val.
    iNext. asimpl.
    rewrite !embed_list_closed; auto.
    rewrite stack_to_list_subst list_to_stack_subst.
    asimpl.
    iApply (clwp_bind [AppRCtx (RecV _)]).
    iApply clwp_wand_l; iSplitR;
      last by iApply (clwp_list_to_stack _ _ []); eauto.
    iIntros (w) "Hst". rewrite app_nil_r.
    iApply clwp_rec; eauto using to_of_val.
    asimpl. rewrite stack_to_list_subst /=.
    iNext.
    iApply clwp_wand_l; iSplitR; last by iApply (clwp_stack_to_list); eauto.
    iIntros (v) "[% Hv]"; subst. done.
  Qed.

  Typeclasses Opaque list_rev.
  Global Opaque list_rev.
  Typeclasses Opaque list_rev_val.
  Global Opaque list_rev_val.

End list_rev.