LogrelCC.fundamental_unary

From LogrelCC Require Export lang rules_unary logrel_unary typing.
From iris.base_logic Require Export invariants.
From iris.proofmode Require Import tactics.

Definition log_typed `{heapG Σ} (Γ : list type) (e : expr) (τ : type) := Δ vs,
  env_Persistent Δ
   Γ ⟧* Δ vs τ ⟧ₑ Δ e.[env_subst vs].
Notation "Γ ⊨ e : τ" := (log_typed Γ e τ) (at level 74, e, τ at next level).

Section typed_interp.
  Context `{heapG Σ}.
  Notation D := (valC -n> iProp Σ).

  (* Put all quantifiers at the outer level *)
  Lemma bin_log_related_alt {Γ e τ} : Γ e : τ Δ vvs K,
    env_Persistent Δ Γ ⟧* Δ vvs interp_ectx (interp τ) K Δ
     WP fill K e.[env_subst (vvs)] {{ _, True}}.
  Proof.
    iIntros (Hlog Δ vvs K ?) "[#HΓ HK]".
    by iApply (Hlog with "HΓ").
  Qed.

   Notation "`' H" := (bin_log_related_alt H) (at level 8).

  Local Tactic Notation "smart_bind"
        constr(HP) uconstr(ctx) ident(v) constr(Hv) :=
    iApply (`'HP _ _ ctx); iFrame "#"; iAlways; iIntros (v) Hv.

  Theorem fundamental Γ e τ : Γ ⊢ₜ e : τ Γ e : τ.
  Proof.
    induction 1; iIntros (Δ vs ) "#HΓ /=".
    - (* var *)
      iDestruct (interp_env_Some_l with "HΓ") as (v) "[% ?]"; first done.
      rewrite /env_subst. simplify_option_eq.
      by iIntros (K) "#HK"; iApply "HK".
    - (* unit *) by iIntros (K) "#HK"; iApply ("HK" $! UnitV).
    - (* nat *) by iIntros (K) "#HK"; iApply ("HK" $! (#nv n)); eauto.
    - (* bool *) by iIntros (K) "#HK"; iApply ("HK" $! (#♭v b)); eauto.
    - (* nat bin op *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 ((BinOpLCtx _ e2.[env_subst vs]) :: K) v "#Hv".
      iDestruct "Hv" as (n1) "%"; subst.
      smart_bind IHtyped2 ((BinOpRCtx _ (#nv n1)) :: K) w "#Hw".
      iDestruct "Hw" as (n2) "%"; subst.
      assert (exists z, binop_eval op (#nv n1) (#nv n2) = Some z) as [? Hop].
      { destruct op; simpl; eauto. }
      iApply wp_bin_op; eauto using to_of_val.
      iApply "HK".
      iNext; destruct op; iPureIntro; simpl in *; inversion Hop; eauto.
    - (* pair *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 ((PairLCtx e2.[env_subst vs]) :: K) v "#Hv".
      smart_bind IHtyped2 ((PairRCtx v) :: K) v' "#Hv'".
      iApply ("HK" $! (PairV _ _)); eauto.
    - (* fst *)
      iIntros (K) "#HK".
      smart_bind IHtyped ((FstCtx) :: K) v "# Hv"; cbn.
      iDestruct "Hv" as (w1 w2) "#[% [H2 H3]]"; subst; simpl.
      iApply wp_fst; eauto.
      iApply "HK"; eauto.
    - (* snd *)
      iIntros (K) "#HK".
      smart_bind IHtyped ((SndCtx) :: K) v "# Hv"; cbn.
      iDestruct "Hv" as (w1 w2) "#[% [H2 H3]]"; subst.
      iApply wp_snd; eauto using to_of_val.
      iApply "HK"; eauto.
    - (* injl *)
      iIntros (K) "#HK".
      smart_bind IHtyped (InjLCtx :: K) v "# Hv".
      iApply ("HK" $! (InjLV _)); eauto.
    - (* injr *)
      iIntros (K) "#HK".
      smart_bind IHtyped (InjRCtx :: K) v "# Hv".
      iApply ("HK" $! (InjRV _)); eauto.
    - (* case *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 ((CaseCtx _ _) :: K) v "# Hv".
      iDestruct (interp_env_length with "HΓ") as %?.
      iDestruct "Hv" as "[Hv|Hv]"; iDestruct "Hv" as (w) "[% Hw]";
        simplify_eq/=.
      + iApply wp_case_injl; eauto using to_of_val, mk_is_Some; asimpl. iNext.
        erewrite typed_subst_head_simpl by naive_solver.
        iApply (IHtyped2 Δ (w :: vs)); auto. iApply interp_env_cons; auto.
      + iApply wp_case_injr; eauto using to_of_val, mk_is_Some; asimpl. iNext.
        erewrite typed_subst_head_simpl by naive_solver.
        iApply (IHtyped3 Δ (w :: vs)); auto. iApply interp_env_cons; auto.
    - (* If *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 (IfCtx _ _ :: K) v "#Hv"; cbn.
      iDestruct "Hv" as ([]) "%"; subst; simpl;
        [iApply wp_if_true| iApply wp_if_false]; iNext;
      [iApply IHtyped2| iApply IHtyped3]; auto.
    - (* Rec *)
      iIntros (K) "#HK".
      iApply ("HK" $! (RecV _)); iClear (K) "HK".
      iAlways. iLöb as "IH". iIntros (w) "#Hw". iIntros (K) "#HK".
      iDestruct (interp_env_length with "HΓ") as %?.
      iApply wp_rec; eauto using to_of_val, mk_is_Some. iNext.
      asimpl. change (Rec _) with (of_val (RecV e.[upn 2 (env_subst vs)])) at 2.
      erewrite typed_subst_head_simpl_2 by naive_solver.
      iApply (IHtyped Δ (_ :: w :: vs)); auto.
      iApply interp_env_cons; iSplit; [|iApply interp_env_cons]; auto.
    - (* Lam *)
      iIntros (K) "#HK".
      iApply ("HK" $! (LamV _)); iClear (K) "HK".
      iAlways. iLöb as "IH". iIntros (w) "#Hw". iIntros (K) "#HK".
      iDestruct (interp_env_length with "HΓ") as %?.
      iApply wp_Lam; eauto using to_of_val, mk_is_Some. iNext.
      asimpl.
      erewrite typed_subst_head_simpl by naive_solver.
      iApply (IHtyped Δ (w :: vs)); auto.
      iApply interp_env_cons; iSplit; auto.
    - (* LetIn *)
      iIntros (K) "#HK".
      iDestruct (interp_env_length with "HΓ") as %?.
      smart_bind IHtyped1 (LetInCtx (e'.[up (env_subst vs)]) :: K) v "#Hv".
      iApply wp_LetIn; eauto.
      iNext; simpl. asimpl.
      erewrite typed_subst_head_simpl by naive_solver.
      iApply (IHtyped2 Δ (v :: vs)); auto.
      iApply interp_env_cons; iSplit; auto.
    - (* LetIn *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 (SeqCtx (e'.[env_subst vs]) :: K) v "#Hv".
      iApply wp_Seq; eauto.
      iNext; simpl. asimpl.
      iApply (IHtyped2 Δ vs); auto.
    - (* app *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 (AppLCtx (e2.[env_subst vs]) :: K) v "#Hv".
      smart_bind IHtyped2 (AppRCtx v :: K) w "#Hw".
      iApply wp_mono; [|iApply "Hv"]; auto.
    - (* TLam *)
      iIntros (K) "#HK".
      iApply ("HK" $! (TLamV _)); iClear (K) "HK".
      iAlways; iIntrosi) "%". iIntros (K) "#HK".
      iApply wp_tapp. iNext.
      iApply (IHtypedi :: Δ)); auto. by iApply interp_env_ren.
    - (* TApp *)
      iIntros (K) "#HK".
      smart_bind IHtyped (TAppCtx :: K) v "#Hv".
      iApply ("Hv" $! (interp τ' Δ)); first by iPureIntro; typeclasses eauto.
      iAlways. iIntros (?) "?"; iApply "HK".
        by iApply interp_subst.
    - (* Fold *)
      iIntros (K) "#HK".
      smart_bind IHtyped (FoldCtx :: K) v "#Hv".
      iApply ("HK" $! (FoldV _)).
      rewrite -(interp_subst _ _ _ _).
      rewrite (fixpoint_unfold (interp_rec1 _ _) _); simpl.
      iAlways; eauto.
    - (* Unfold *)
      iIntros (K) "#HK".
      smart_bind IHtyped (UnfoldCtx :: K) v "#Hv".
      rewrite /= (fixpoint_unfold (interp_rec1 _ _) _); simpl.
      change (fixpoint _) with ( TRec τ Δ); simpl.
      iDestruct "Hv" as (w) "#[% Hw]"; subst.
      iApply wp_unfold; cbn; auto using to_of_val.
      iNext; iApply "HK". by iApply interp_subst.
    - (* Fork *)
      iIntros (K) "#HK".
      iApply wp_fork; iSplitL; first by iApply ("HK" $! UnitV).
      iApply (`'IHtyped _ _ []); iFrame "#"; simpl. iNext.
      iAlways. iIntros (w) "%"; simplify_eq; simpl. by iApply wp_value.
    - (* Alloc *)
      iIntros (K) "#HK".
      smart_bind IHtyped (AllocCtx :: K) v "#Hv".
      iApply wp_alloc; auto using to_of_val.
      iNext; iIntros (l) "Hl".
      iMod (inv_alloc (logN.@l) _ ( w : val, l ↦ᵢ w τ Δ w)%I with "[Hl]")
        as "HN".
      { iNext; iExists _; eauto. }
      iApply ("HK" $! (LocV _)).
      iExists _; iSplit; eauto.
    - (* Load *)
      iIntros (K) "#HK".
      smart_bind IHtyped (LoadCtx :: K) v "#Hv"; simpl.
      iDestruct "Hv" as (l) "[% #Hv]"; subst.
      iApply wp_atomic_under_ectx; eauto.
      iInv (logN .@ l) as (w) "[Hw1 #Hw2]" "Hclose".
      iModIntro.
      iApply (wp_load [] with "[-]"); iFrame.
      iNext. iIntros "Hw1". iApply wp_value; eauto using to_of_val.
      iMod ("Hclose" with "[Hw1]") as "_"; first eauto.
      iModIntro. by iApply "HK".
    - (* Store *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 (StoreLCtx _ :: K) v "#Hv"; simpl.
      smart_bind IHtyped2 (StoreRCtx _ :: K) v' "#Hv'"; simpl.
      iDestruct "Hv" as (l) "[% #Hv]"; subst.
      iApply wp_atomic_under_ectx; eauto.
      iInv (logN .@ l) as (w) "[Hw1 #Hw2]" "Hclose".
      iModIntro. simpl.
      iApply (wp_store [] with "[-]"); eauto using to_of_val; iFrame.
      iNext. iIntros "Hw1". iApply wp_value; eauto using to_of_val.
      iMod ("Hclose" with "[Hw1]") as "_"; first eauto.
      iModIntro. by iApply "HK".
    - (* CAS *)
      iIntros (K) "#HK".
      smart_bind IHtyped1 (CasLCtx _ _ :: K) v "#Hv"; simpl.
      smart_bind IHtyped2 (CasMCtx _ _ :: K) v' "#Hv'"; simpl.
      smart_bind IHtyped3 (CasRCtx _ _ :: K) v'' "#Hv''"; simpl.
      iDestruct "Hv" as (l) "[% #Hv]"; subst.
      iApply wp_atomic_under_ectx; eauto.
      iInv (logN .@ l) as (w) "[Hw1 #Hw2]" "Hclose".
      iModIntro. simpl.
      destruct (decide (v' = w)) as [|Hneq]; subst.
      + iApply (wp_cas_suc [] with "[-]"); eauto using to_of_val; iFrame.
        iNext. iIntros "Hw1". iApply wp_value; eauto using to_of_val.
        iMod ("Hclose" with "[Hw1]") as "_"; first eauto.
        iModIntro. iApply ("HK" $! (#♭v true)); eauto.
      + iApply (wp_cas_fail [] with "[-]"); eauto using to_of_val; iFrame.
        iNext. iIntros "Hw1". iApply wp_value; eauto using to_of_val.
        iMod ("Hclose" with "[Hw1]") as "_"; first eauto.
        iModIntro. iApply ("HK" $! (#♭v false)); eauto.
    - (* Callcc *)
      iIntros (K) "#HK".
      iApply wp_callcc.
      iNext.
      iDestruct (interp_env_length with "HΓ") as %?.
      asimpl.
      rewrite (typed_subst_head_simpl (TCont τ :: Γ) τ e (ContV _)); trivial;
        last by naive_solver.
      iApply (IHtyped Δ (ContV K :: vs) with "[]"); auto.
      rewrite interp_env_cons; simpl; eauto.
    - iIntros (K) "#HK".
      smart_bind IHtyped1 (ThrowLCtx _ :: K) v "#Hv"; simpl.
      smart_bind IHtyped2 (ThrowRCtx _ :: K) v' "#Hv'"; simpl.
      iDestruct "Hv'" as (K') "[% #Hv']"; subst.
      iApply wp_throw; eauto using to_of_val.
      iNext. by iApply "Hv'".
  Qed.
End typed_interp.