LogrelCC.fundamental_binary
From LogrelCC Require Export logrel_binary.
From iris.proofmode Require Import tactics.
From iris.algebra Require Export list.
Section bin_log_def.
Context `{heapG Σ, cfgSG Σ}.
Notation D := (prodC valC valC -n> iProp Σ).
Definition bin_log_related (Γ : list type) (e e' : expr) (τ : type) := ∀ Δ vvs ρ,
env_Persistent Δ →
spec_ctx ρ ∧
⟦ Γ ⟧* Δ vvs ⊢ ⟦ τ ⟧ₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]).
End bin_log_def.
Notation "Γ ⊨ e '≤log≤' e' : τ" :=
(bin_log_related Γ e e' τ) (at level 74, e, e', τ at next level).
Section fundamental.
Context `{heapG Σ, cfgSG Σ}.
Notation D := (prodC valC valC -n> iProp Σ).
Implicit Types e : expr.
Implicit Types Δ : listC D.
Hint Resolve to_of_val.
(* Put all quantifiers at the outer level *)
Lemma bin_log_related_alt {Γ e e' τ} : Γ ⊨ e ≤log≤ e' : τ → ∀ Δ vvs ρ j KK,
env_Persistent Δ →
spec_ctx ρ ∗ ⟦ Γ ⟧* Δ vvs ∗ interp_ectx (interp τ) KK Δ ∗
j ⤇ fill (KK.2) (e'.[env_subst (vvs.*2)])
⊢ WP fill (KK.1) e.[env_subst (vvs.*1)] {{ v, ∃ v',
j ⤇ of_val v' }}.
Proof.
iIntros (Hlog Δ vvs ρ j K ?) "(#Hs & HΓ & HKK & Hj)".
iApply (Hlog with "[HΓ]"); iFrame; eauto.
Qed.
Notation "`' H" := (bin_log_related_alt H) (at level 8).
Local Tactic Notation "smart_bind"
constr(HP) uconstr(ctx) ident(vv) ident(j) constr(Hv) :=
iApply (`'HP _ _ _ _ ctx); iFrame "#"; iFrame; iAlways; iIntros (vv j) Hv.
Lemma bin_log_related_var Γ x τ :
Γ !! x = Some τ → Γ ⊨ Var x ≤log≤ Var x : τ.
Proof.
iIntros (? Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_Some_l with "HΓ") as ([v v']) "[% ?]"; first done.
rewrite /env_subst !list_lookup_fmap; simplify_option_eq.
by iApply ("HKK" $! (v,v') j with "[-]"); iFrame.
Qed.
Lemma bin_log_related_unit Γ : Γ ⊨ Unit ≤log≤ Unit : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
by iApply ("HKK" $! (UnitV, UnitV)); iFrame.
Qed.
Lemma bin_log_related_nat Γ n : Γ ⊨ #n n ≤log≤ #n n : TNat.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (#nv _, #nv _)); iFrame; eauto.
Qed.
Lemma bin_log_related_bool Γ b : Γ ⊨ #♭ b ≤log≤ #♭ b : TBool.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (#♭v _, #♭v _)); iFrame; eauto.
Qed.
Lemma bin_log_related_pair Γ e1 e2 e1' e2' τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ Pair e1 e2 ≤log≤ Pair e1' e2' : TProd τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (PairLCtx _ :: KK.1, PairLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]".
smart_bind IHHtyped2 (PairRCtx _ :: KK.1, PairRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]".
iApply ("HKK" $! (PairV _ _, PairV _ _)); iFrame; eauto.
Qed.
Lemma bin_log_related_fst Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : TProd τ1 τ2) :
Γ ⊨ Fst e ≤log≤ Fst e' : τ1.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (FstCtx :: KK.1, FstCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
iApply wp_fst; eauto. iNext.
iMod (step_fst with "[Hj1]") as "Hw"; eauto.
iApply "HKK"; iFrame; eauto.
Qed.
Lemma bin_log_related_snd Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : TProd τ1 τ2) :
Γ ⊨ Snd e ≤log≤ Snd e' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (SndCtx :: KK.1, SndCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
iApply wp_snd; eauto. iNext.
iMod (step_snd with "[Hj1]") as "Hw"; eauto.
iApply "HKK"; iFrame; eauto.
Qed.
Lemma bin_log_related_injl Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ1) :
Γ ⊨ InjL e ≤log≤ InjL e' : (TSum τ1 τ2).
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (InjLCtx :: KK.1, InjLCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iApply ("HKK" $! (InjLV _, InjLV _)); iFrame; eauto.
Qed.
Lemma bin_log_related_injr Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ InjR e ≤log≤ InjR e' : TSum τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (InjRCtx :: KK.1, InjRCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iApply ("HKK" $! (InjRV _, InjRV _)); iFrame; eauto.
Qed.
Lemma bin_log_related_case Γ e0 e1 e2 e0' e1' e2' τ1 τ2 τ3
(Hclosed2 : ∀ f, e1.[upn (S (length Γ)) f] = e1)
(Hclosed3 : ∀ f, e2.[upn (S (length Γ)) f] = e2)
(Hclosed2' : ∀ f, e1'.[upn (S (length Γ)) f] = e1')
(Hclosed3' : ∀ f, e2'.[upn (S (length Γ)) f] = e2')
(IHHtyped1 : Γ ⊨ e0 ≤log≤ e0' : TSum τ1 τ2)
(IHHtyped2 : τ1 :: Γ ⊨ e1 ≤log≤ e1' : τ3)
(IHHtyped3 : τ2 :: Γ ⊨ e2 ≤log≤ e2' : τ3) :
Γ ⊨ Case e0 e1 e2 ≤log≤ Case e0' e1' e2' : τ3.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
smart_bind IHHtyped1 (CaseCtx _ _ :: KK.1, CaseCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as "[Hvv|Hvv]";
iDestruct "Hvv" as ([w w']) "[% Hw]"; simplify_eq; simpl.
- iMod (step_case_inl with "[Hs Hj1]") as "Hz"; eauto.
iApply wp_case_injl; eauto using to_of_val. iNext.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped2 _ ((w,w') :: vvs)). repeat iSplit; eauto.
iApply interp_env_cons; auto.
- iMod (step_case_inr with "[Hs Hj1]") as "Hz"; eauto.
iApply wp_case_injr; eauto using to_of_val. iNext.
asimpl. erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped3 _ ((w,w') :: vvs)); repeat iSplit; eauto.
iApply interp_env_cons; auto.
Qed.
Lemma bin_log_related_if Γ e0 e1 e2 e0' e1' e2' τ
(IHHtyped1 : Γ ⊨ e0 ≤log≤ e0' : TBool)
(IHHtyped2 : Γ ⊨ e1 ≤log≤ e1' : τ)
(IHHtyped3 : Γ ⊨ e2 ≤log≤ e2' : τ) :
Γ ⊨ If e0 e1 e2 ≤log≤ If e0' e1' e2' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (IfCtx _ _ :: KK.1, IfCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
destruct vv as [v1 v2].
iDestruct "Hvv" as ([]) "[% %]"; simplify_eq/=.
- iMod (step_if_true with "[-]") as "Hz"; eauto.
iApply wp_if_true. iNext. iApply `'IHHtyped2; eauto.
- iMod (step_if_false with "[-]") as "Hz"; eauto.
iApply wp_if_false. iNext. iApply `'IHHtyped3; eauto.
Qed.
Lemma bin_log_related_nat_binop Γ op e1 e2 e1' e2'
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : TNat)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : TNat) :
Γ ⊨ BinOp op e1 e2 ≤log≤ BinOp op e1' e2' : binop_res_type op.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (BinOpLCtx _ _ :: KK.1, BinOpLCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (BinOpRCtx _ _ :: KK.1, BinOpRCtx _ _ :: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
destruct vv as [v1 v2]; destruct ww as [w1 w2].
iDestruct "Hvv" as (n) "[% %]"; simplify_eq/=.
iDestruct "Hww" as (n') "[% %]"; simplify_eq/=.
destruct op;
iMod (step_nat_binop _ _ _ _ _ _ _ (#nv _) (#nv _) with "[-]") as "Hz";
eauto; simpl; eauto;
iApply wp_bin_op; eauto; iNext;
try iApply ("HKK" $! (#nv _, #nv _)); eauto;
try iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
Qed.
Lemma bin_log_related_rec Γ (e e' : expr) τ1 τ2
(Hclosed : ∀ f, e.[upn (S (S (length Γ))) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (S (length Γ))) f] = e')
(IHHtyped : TArrow τ1 τ2 :: τ1 :: Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ Rec e ≤log≤ Rec e' : TArrow τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (RecV _, RecV _)); iFrame; iAlways. iClear "HKK".
iLöb as "IH".
iIntros ([v v']) "#Hvv". iIntros (j' K') "[Hj' #HKK]".
iDestruct (interp_env_length with "HΓ") as %?.
iApply wp_rec; eauto using to_of_val. iNext.
iMod (step_rec with "[$Hj']") as "Hz"; eauto.
asimpl. change (Rec ?e) with (of_val (RecV e)).
erewrite !n_closed_subst_head_simpl_2 by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((_,_) :: (v,v') :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_Lam Γ (e e' : expr) τ1 τ2
(Hclosed : ∀ f, e.[upn (S (length Γ)) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (length Γ)) f] = e')
(IHHtyped : τ1 :: Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ Lam e ≤log≤ Lam e' : TArrow τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (LamV _, LamV _)); iFrame; iAlways. iClear "HKK".
iLöb as "IH".
iIntros ([v v']) "#Hvv". iIntros (j' K') "[Hj' #HKK]".
iDestruct (interp_env_length with "HΓ") as %?.
iApply wp_Lam; eauto. iNext.
iMod (step_Lam with "[$Hj']") as "Hz"; eauto.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((v,v') :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_LetIn Γ (e1 e1' e2 e2' : expr) τ1 τ2
(Hclosed2 : ∀ f, e2.[upn (S (length Γ)) f] = e2)
(Hclosed2' : ∀ f, e2'.[upn (S (length Γ)) f] = e2')
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : τ1 :: Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ LetIn e1 e2 ≤log≤ LetIn e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
smart_bind IHHtyped1 (LetInCtx _ :: KK.1, LetInCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iApply wp_LetIn; eauto. iNext.
iMod (step_LetIn with "[-]") as "Hz"; eauto.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped2 _ (vv :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_Seq Γ (e1 e1' e2 e2' : expr) τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ Seq e1 e2 ≤log≤ Seq e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (SeqCtx _ :: KK.1, SeqCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iApply wp_Seq; eauto. iNext.
iMod (step_Seq with "[-]") as "Hz"; eauto.
asimpl.
iApply (`'IHHtyped2 _ vvs); eauto.
Qed.
Lemma bin_log_related_app Γ e1 e2 e1' e2' τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : TArrow τ1 τ2)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ1) :
Γ ⊨ App e1 e2 ≤log≤ App e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (AppLCtx _ :: KK.1, AppLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
smart_bind IHHtyped2 (AppRCtx _ :: KK.1, AppRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]".
iApply "Hvv"; auto.
Qed.
Lemma bin_log_related_tlam Γ e e' τ
(IHHtyped : (subst (ren (+1)) <$> Γ) ⊨ e ≤log≤ e' : τ) :
Γ ⊨ TLam e ≤log≤ TLam e' : TForall τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (TLamV _, TLamV _)); iFrame; iAlways. iClear "HKK".
iIntros (τi ? KK' j') "[Hv #HKK] /=".
iApply wp_tapp; iNext.
iMod (step_tlam with "[-]") as "Hz"; eauto.
iApply (`'IHHtyped (τi :: Δ)); iFrame "#"; iFrame.
by rewrite interp_env_ren.
Qed.
Lemma bin_log_related_tapp Γ e e' τ τ'
(IHHtyped : Γ ⊨ e ≤log≤ e' : TForall τ) :
Γ ⊨ TApp e ≤log≤ TApp e' : τ.[τ'/].
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (TAppCtx :: KK.1, TAppCtx :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iSpecialize ("Hvv" $! (interp τ' Δ) with "[#]"); [iPureIntro; apply _|].
iApply "Hvv"; iFrame; eauto.
iAlways. iIntros (??) "?". iApply "HKK".
erewrite <- (interp_subst _ _ _ _); eauto.
Qed.
Lemma bin_log_related_fold Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ.[(TRec τ)/]) :
Γ ⊨ Fold e ≤log≤ Fold e' : TRec τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (FoldCtx :: KK.1, FoldCtx :: KK.2)
vv j1 "[Hj1 #Hvv]".
destruct vv as [v1 v2].
iApply ("HKK" $! (FoldV _, FoldV _)); iFrame; simpl.
rewrite -> (fixpoint_unfold (interp_rec1 _ _) _).
rewrite /= -(interp_subst _ _ _ _).
iAlways; iExists (_, _); eauto.
Qed.
Lemma bin_log_related_unfold Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : TRec τ) :
Γ ⊨ Unfold e ≤log≤ Unfold e' : τ.[(TRec τ)/].
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (UnfoldCtx :: KK.1, UnfoldCtx :: KK.2)
vv j1 "[Hj1 #Hvv]".
destruct vv as [v1 v2].
rewrite /= (fixpoint_unfold (interp_rec1 _ _) _) /=.
change (fixpoint _) with (interp (TRec τ) Δ).
iDestruct "Hvv" as ([w w']) "#[% Hiz]"; simplify_eq/=.
iMod (step_Fold with "[-]") as "Hz"; eauto.
iApply wp_unfold; cbn; auto.
iNext.
iApply ("HKK" $! (w, w')); iFrame.
by rewrite <- (interp_subst _ _ _ _).
Qed.
Lemma bin_log_related_fork Γ e e'
(IHHtyped : Γ ⊨ e ≤log≤ e' : TUnit) :
Γ ⊨ Fork e ≤log≤ Fork e' : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iMod (step_fork with "[-]") as (j') "[Hj Hj']"; eauto.
iApply wp_fork; iSplitL "Hj"; iNext.
- iApply ("HKK" $! (UnitV, UnitV)); eauto.
- iApply wp_wand_l; iSplitR; [|iApply (`'IHHtyped _ _ _ _ ([], []))]; eauto.
iFrame; iFrame "#".
iAlways; iIntros ([]?) "[? [% %]]"; simpl in *; subst; simpl.
iApply wp_value; eauto. iExists UnitV; eauto.
Qed.
Lemma bin_log_related_alloc Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ) :
Γ ⊨ Alloc e ≤log≤ Alloc e' : Tref τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (AllocCtx :: KK.1, AllocCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iMod (step_alloc with "[-]") as (l') "[Hj Hl']"; eauto.
iApply wp_alloc; eauto.
iNext. iIntros (l) "Hl".
iMod (inv_alloc (logN .@ (l,l')) _ (∃ w : val * val,
l ↦ᵢ w.1 ∗ l' ↦ₛ w.2 ∗ interp τ Δ w)%I with "[Hl Hl']") as "HN"; eauto.
{ iNext. iExists vv; iFrame; iFrame "#". }
iApply ("HKK" $! (LocV _, LocV _)); iFrame.
iExists (_, _); eauto.
Qed.
Lemma bin_log_related_load Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : (Tref τ)) :
Γ ⊨ Load e ≤log≤ Load e' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (LoadCtx :: KK.1, LoadCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([w w']) "[Hw1 [>Hw2 #Hw]]" "Hclose"; simpl.
iModIntro.
iApply (wp_load' with "[-]"); iFrame.
iNext. iIntros "Hw1".
iMod (step_load with "[Hj1 Hw2]") as "[Hj1 Hw2]";
[solve_ndisj|by iFrame|].
iApply wp_value; eauto.
iMod ("Hclose" with "[Hw1 Hw2]") as "_".
{ iNext. iExists (w,w'); by iFrame. }
iModIntro. iApply ("HKK" $! (w, w')); eauto.
Qed.
Lemma bin_log_related_store Γ e1 e2 e1' e2' τ
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : (Tref τ))
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ) :
Γ ⊨ Store e1 e2 ≤log≤ Store e1' e2' : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (StoreLCtx _ :: KK.1, StoreLCtx _:: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (StoreRCtx _ :: KK.1, StoreRCtx _:: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([v v']) "[Hv1 [>Hv2 #Hv]]" "Hclose".
iModIntro.
iApply (wp_store' with "[-]"); auto using to_of_val.
iFrame. iNext. iIntros "Hv1".
iMod (step_store with "[$Hs Hj2 Hv2]") as "[Hj2 Hv2]"; eauto;
[solve_ndisj | by iFrame|].
iApply wp_value; eauto.
iMod ("Hclose" with "[Hv1 Hv2]") as "_".
{ iNext; iExists _; simpl; iFrame; iFrame "#". }
iModIntro. iApply ("HKK" $! (UnitV, UnitV)); eauto.
Qed.
Lemma bin_log_related_CAS Γ e1 e2 e3 e1' e2' e3' τ
(HEqτ : EqType τ)
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : Tref τ)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ)
(IHHtyped3 : Γ ⊨ e3 ≤log≤ e3' : τ) :
Γ ⊨ CAS e1 e2 e3 ≤log≤ CAS e1' e2' e3' : TBool.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (CasLCtx _ _ :: KK.1, CasLCtx _ _:: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (CasMCtx _ _ :: KK.1, CasMCtx _ _:: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
smart_bind IHHtyped3 (CasRCtx _ _ :: KK.1, CasRCtx _ _:: KK.2)
uu j3 "[Hj3 #Huu]"; simpl.
destruct ww as [w w']; destruct uu as [u u'].
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([v v']) "[Hv1 [>Hv2 #Hv]]" "Hclose".
iModIntro.
iPoseProof ("Hv") as "Hv'".
rewrite {2}[⟦ τ ⟧ Δ (v, v')]interp_EqType_agree; trivial.
iMod "Hv'" as %Hv'; subst.
destruct (decide (v' = w)) as [|Hneq]; subst.
- iAssert (▷ ⌜w' = w⌝)%I as ">%".
{ rewrite ?interp_EqType_agree; trivial. by iSimplifyEq. }
simpl. iApply (wp_cas_suc' with "[-]"); eauto using to_of_val.
iFrame. iNext. iIntros "Hv1".
iMod (step_cas_suc
with "[Hj3 Hv2]") as "[Hj3 Hv2]"; simpl; eauto; first solve_ndisj.
iFrame. iFrame "Hs".
iApply wp_value; eauto.
iMod ("Hclose" with "[Hv1 Hv2]").
{ iNext; iExists (_, _); by iFrame. }
iModIntro. iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
- iAssert (▷ ⌜v' ≠ w'⌝)%I as ">%".
{ rewrite ?interp_EqType_agree; trivial. iSimplifyEq. auto. }
simpl. iApply (wp_cas_fail' with "[-]"); eauto.
iFrame. iNext. iIntros "Hv1".
iApply wp_value; eauto.
iMod (step_cas_fail
with "[$Hs Hj3 Hv2]") as "[Hj3 Hv2]"; simpl; eauto; first solve_ndisj.
iFrame.
iMod ("Hclose" with "[Hv1 Hv2]").
{ iNext; iExists (_, _); by iFrame. }
iModIntro. iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
Qed.
Lemma bin_log_related_callcc Γ e e' τ
(Hclosed : ∀ f, e.[upn (S (length Γ)) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (length Γ)) f] = e')
(IHHtyped : TCont τ :: Γ ⊨ e ≤log≤ e' : τ) :
Γ ⊨ Callcc e ≤log≤ Callcc e' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
iMod (step_callcc with "[Hj]") as "Hj"; eauto.
iApply wp_callcc; iNext.
asimpl.
change (Cont (KK.2)) with (of_val (ContV (KK.2))).
change (Cont (KK.1)) with (of_val (ContV (KK.1))).
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((ContV (KK.1), ContV (KK.2)) :: vvs)).
iFrame; iFrame "#".
iApply interp_env_cons; iFrame "#"; auto.
simpl; iExists _, _; eauto.
Qed.
Lemma bin_log_related_throw Γ e1 e2 e1' e2' τ τ'
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : TCont τ) :
Γ ⊨ Throw e1 e2 ≤log≤ Throw e1' e2' : τ'.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (ThrowLCtx _ :: KK.1, ThrowLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". simpl.
smart_bind IHHtyped2 (ThrowRCtx _ :: KK.1, ThrowRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
destruct ww as [w1 w2].
iDestruct "Hww" as (K K') "[% #HKK2]"; simplify_eq; simpl.
iMod (step_throw with "[Hj2]") as "Hj2"; eauto.
iApply wp_throw; eauto.
iNext.
iApply "HKK2"; eauto.
Qed.
Theorem binary_fundamental Γ e τ :
Γ ⊢ₜ e : τ → Γ ⊨ e ≤log≤ e : τ.
Proof.
induction 1.
- by apply bin_log_related_var.
- by apply bin_log_related_unit.
- by apply bin_log_related_nat.
- by apply bin_log_related_bool.
- apply bin_log_related_nat_binop; eauto.
- apply bin_log_related_pair; eauto.
- eapply bin_log_related_fst; eauto.
- eapply bin_log_related_snd; eauto.
- eapply bin_log_related_injl; eauto.
- eapply bin_log_related_injr; eauto.
- eapply bin_log_related_case; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_if; eauto.
- eapply bin_log_related_rec; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_Lam; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_LetIn; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_Seq; eauto.
- eapply (bin_log_related_app _ _ _ _ _ τ1); auto.
- eapply bin_log_related_tlam; eauto with typeclass_instances.
- eapply bin_log_related_tapp; eauto.
- eapply bin_log_related_fold; eauto.
- eapply bin_log_related_unfold; eauto.
- eapply bin_log_related_fork; eauto.
- eapply bin_log_related_alloc; eauto.
- eapply bin_log_related_load; eauto.
- eapply (bin_log_related_store _ _ _ _ _ τ) ; eauto.
- eapply bin_log_related_CAS; eauto.
- eapply bin_log_related_callcc; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_throw; eauto.
Qed.
End fundamental.
From iris.proofmode Require Import tactics.
From iris.algebra Require Export list.
Section bin_log_def.
Context `{heapG Σ, cfgSG Σ}.
Notation D := (prodC valC valC -n> iProp Σ).
Definition bin_log_related (Γ : list type) (e e' : expr) (τ : type) := ∀ Δ vvs ρ,
env_Persistent Δ →
spec_ctx ρ ∧
⟦ Γ ⟧* Δ vvs ⊢ ⟦ τ ⟧ₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]).
End bin_log_def.
Notation "Γ ⊨ e '≤log≤' e' : τ" :=
(bin_log_related Γ e e' τ) (at level 74, e, e', τ at next level).
Section fundamental.
Context `{heapG Σ, cfgSG Σ}.
Notation D := (prodC valC valC -n> iProp Σ).
Implicit Types e : expr.
Implicit Types Δ : listC D.
Hint Resolve to_of_val.
(* Put all quantifiers at the outer level *)
Lemma bin_log_related_alt {Γ e e' τ} : Γ ⊨ e ≤log≤ e' : τ → ∀ Δ vvs ρ j KK,
env_Persistent Δ →
spec_ctx ρ ∗ ⟦ Γ ⟧* Δ vvs ∗ interp_ectx (interp τ) KK Δ ∗
j ⤇ fill (KK.2) (e'.[env_subst (vvs.*2)])
⊢ WP fill (KK.1) e.[env_subst (vvs.*1)] {{ v, ∃ v',
j ⤇ of_val v' }}.
Proof.
iIntros (Hlog Δ vvs ρ j K ?) "(#Hs & HΓ & HKK & Hj)".
iApply (Hlog with "[HΓ]"); iFrame; eauto.
Qed.
Notation "`' H" := (bin_log_related_alt H) (at level 8).
Local Tactic Notation "smart_bind"
constr(HP) uconstr(ctx) ident(vv) ident(j) constr(Hv) :=
iApply (`'HP _ _ _ _ ctx); iFrame "#"; iFrame; iAlways; iIntros (vv j) Hv.
Lemma bin_log_related_var Γ x τ :
Γ !! x = Some τ → Γ ⊨ Var x ≤log≤ Var x : τ.
Proof.
iIntros (? Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_Some_l with "HΓ") as ([v v']) "[% ?]"; first done.
rewrite /env_subst !list_lookup_fmap; simplify_option_eq.
by iApply ("HKK" $! (v,v') j with "[-]"); iFrame.
Qed.
Lemma bin_log_related_unit Γ : Γ ⊨ Unit ≤log≤ Unit : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
by iApply ("HKK" $! (UnitV, UnitV)); iFrame.
Qed.
Lemma bin_log_related_nat Γ n : Γ ⊨ #n n ≤log≤ #n n : TNat.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (#nv _, #nv _)); iFrame; eauto.
Qed.
Lemma bin_log_related_bool Γ b : Γ ⊨ #♭ b ≤log≤ #♭ b : TBool.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (#♭v _, #♭v _)); iFrame; eauto.
Qed.
Lemma bin_log_related_pair Γ e1 e2 e1' e2' τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ Pair e1 e2 ≤log≤ Pair e1' e2' : TProd τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (PairLCtx _ :: KK.1, PairLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]".
smart_bind IHHtyped2 (PairRCtx _ :: KK.1, PairRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]".
iApply ("HKK" $! (PairV _ _, PairV _ _)); iFrame; eauto.
Qed.
Lemma bin_log_related_fst Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : TProd τ1 τ2) :
Γ ⊨ Fst e ≤log≤ Fst e' : τ1.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (FstCtx :: KK.1, FstCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
iApply wp_fst; eauto. iNext.
iMod (step_fst with "[Hj1]") as "Hw"; eauto.
iApply "HKK"; iFrame; eauto.
Qed.
Lemma bin_log_related_snd Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : TProd τ1 τ2) :
Γ ⊨ Snd e ≤log≤ Snd e' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (SndCtx :: KK.1, SndCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([w1 w1'] [w2 w2']) "#[% [Hw1 Hw2]]"; simplify_eq.
iApply wp_snd; eauto. iNext.
iMod (step_snd with "[Hj1]") as "Hw"; eauto.
iApply "HKK"; iFrame; eauto.
Qed.
Lemma bin_log_related_injl Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ1) :
Γ ⊨ InjL e ≤log≤ InjL e' : (TSum τ1 τ2).
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (InjLCtx :: KK.1, InjLCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iApply ("HKK" $! (InjLV _, InjLV _)); iFrame; eauto.
Qed.
Lemma bin_log_related_injr Γ e e' τ1 τ2
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ InjR e ≤log≤ InjR e' : TSum τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (InjRCtx :: KK.1, InjRCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iApply ("HKK" $! (InjRV _, InjRV _)); iFrame; eauto.
Qed.
Lemma bin_log_related_case Γ e0 e1 e2 e0' e1' e2' τ1 τ2 τ3
(Hclosed2 : ∀ f, e1.[upn (S (length Γ)) f] = e1)
(Hclosed3 : ∀ f, e2.[upn (S (length Γ)) f] = e2)
(Hclosed2' : ∀ f, e1'.[upn (S (length Γ)) f] = e1')
(Hclosed3' : ∀ f, e2'.[upn (S (length Γ)) f] = e2')
(IHHtyped1 : Γ ⊨ e0 ≤log≤ e0' : TSum τ1 τ2)
(IHHtyped2 : τ1 :: Γ ⊨ e1 ≤log≤ e1' : τ3)
(IHHtyped3 : τ2 :: Γ ⊨ e2 ≤log≤ e2' : τ3) :
Γ ⊨ Case e0 e1 e2 ≤log≤ Case e0' e1' e2' : τ3.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
smart_bind IHHtyped1 (CaseCtx _ _ :: KK.1, CaseCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as "[Hvv|Hvv]";
iDestruct "Hvv" as ([w w']) "[% Hw]"; simplify_eq; simpl.
- iMod (step_case_inl with "[Hs Hj1]") as "Hz"; eauto.
iApply wp_case_injl; eauto using to_of_val. iNext.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped2 _ ((w,w') :: vvs)). repeat iSplit; eauto.
iApply interp_env_cons; auto.
- iMod (step_case_inr with "[Hs Hj1]") as "Hz"; eauto.
iApply wp_case_injr; eauto using to_of_val. iNext.
asimpl. erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped3 _ ((w,w') :: vvs)); repeat iSplit; eauto.
iApply interp_env_cons; auto.
Qed.
Lemma bin_log_related_if Γ e0 e1 e2 e0' e1' e2' τ
(IHHtyped1 : Γ ⊨ e0 ≤log≤ e0' : TBool)
(IHHtyped2 : Γ ⊨ e1 ≤log≤ e1' : τ)
(IHHtyped3 : Γ ⊨ e2 ≤log≤ e2' : τ) :
Γ ⊨ If e0 e1 e2 ≤log≤ If e0' e1' e2' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (IfCtx _ _ :: KK.1, IfCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
destruct vv as [v1 v2].
iDestruct "Hvv" as ([]) "[% %]"; simplify_eq/=.
- iMod (step_if_true with "[-]") as "Hz"; eauto.
iApply wp_if_true. iNext. iApply `'IHHtyped2; eauto.
- iMod (step_if_false with "[-]") as "Hz"; eauto.
iApply wp_if_false. iNext. iApply `'IHHtyped3; eauto.
Qed.
Lemma bin_log_related_nat_binop Γ op e1 e2 e1' e2'
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : TNat)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : TNat) :
Γ ⊨ BinOp op e1 e2 ≤log≤ BinOp op e1' e2' : binop_res_type op.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (BinOpLCtx _ _ :: KK.1, BinOpLCtx _ _ :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (BinOpRCtx _ _ :: KK.1, BinOpRCtx _ _ :: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
destruct vv as [v1 v2]; destruct ww as [w1 w2].
iDestruct "Hvv" as (n) "[% %]"; simplify_eq/=.
iDestruct "Hww" as (n') "[% %]"; simplify_eq/=.
destruct op;
iMod (step_nat_binop _ _ _ _ _ _ _ (#nv _) (#nv _) with "[-]") as "Hz";
eauto; simpl; eauto;
iApply wp_bin_op; eauto; iNext;
try iApply ("HKK" $! (#nv _, #nv _)); eauto;
try iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
Qed.
Lemma bin_log_related_rec Γ (e e' : expr) τ1 τ2
(Hclosed : ∀ f, e.[upn (S (S (length Γ))) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (S (length Γ))) f] = e')
(IHHtyped : TArrow τ1 τ2 :: τ1 :: Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ Rec e ≤log≤ Rec e' : TArrow τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (RecV _, RecV _)); iFrame; iAlways. iClear "HKK".
iLöb as "IH".
iIntros ([v v']) "#Hvv". iIntros (j' K') "[Hj' #HKK]".
iDestruct (interp_env_length with "HΓ") as %?.
iApply wp_rec; eauto using to_of_val. iNext.
iMod (step_rec with "[$Hj']") as "Hz"; eauto.
asimpl. change (Rec ?e) with (of_val (RecV e)).
erewrite !n_closed_subst_head_simpl_2 by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((_,_) :: (v,v') :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_Lam Γ (e e' : expr) τ1 τ2
(Hclosed : ∀ f, e.[upn (S (length Γ)) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (length Γ)) f] = e')
(IHHtyped : τ1 :: Γ ⊨ e ≤log≤ e' : τ2) :
Γ ⊨ Lam e ≤log≤ Lam e' : TArrow τ1 τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (LamV _, LamV _)); iFrame; iAlways. iClear "HKK".
iLöb as "IH".
iIntros ([v v']) "#Hvv". iIntros (j' K') "[Hj' #HKK]".
iDestruct (interp_env_length with "HΓ") as %?.
iApply wp_Lam; eauto. iNext.
iMod (step_Lam with "[$Hj']") as "Hz"; eauto.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((v,v') :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_LetIn Γ (e1 e1' e2 e2' : expr) τ1 τ2
(Hclosed2 : ∀ f, e2.[upn (S (length Γ)) f] = e2)
(Hclosed2' : ∀ f, e2'.[upn (S (length Γ)) f] = e2')
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : τ1 :: Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ LetIn e1 e2 ≤log≤ LetIn e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
smart_bind IHHtyped1 (LetInCtx _ :: KK.1, LetInCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iApply wp_LetIn; eauto. iNext.
iMod (step_LetIn with "[-]") as "Hz"; eauto.
asimpl.
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped2 _ (vv :: vvs)); repeat iSplit; eauto.
rewrite !interp_env_cons; iSplit; auto.
Qed.
Lemma bin_log_related_Seq Γ (e1 e1' e2 e2' : expr) τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ1)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ2) :
Γ ⊨ Seq e1 e2 ≤log≤ Seq e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (SeqCtx _ :: KK.1, SeqCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iApply wp_Seq; eauto. iNext.
iMod (step_Seq with "[-]") as "Hz"; eauto.
asimpl.
iApply (`'IHHtyped2 _ vvs); eauto.
Qed.
Lemma bin_log_related_app Γ e1 e2 e1' e2' τ1 τ2
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : TArrow τ1 τ2)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ1) :
Γ ⊨ App e1 e2 ≤log≤ App e1' e2' : τ2.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (AppLCtx _ :: KK.1, AppLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
smart_bind IHHtyped2 (AppRCtx _ :: KK.1, AppRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]".
iApply "Hvv"; auto.
Qed.
Lemma bin_log_related_tlam Γ e e' τ
(IHHtyped : (subst (ren (+1)) <$> Γ) ⊨ e ≤log≤ e' : τ) :
Γ ⊨ TLam e ≤log≤ TLam e' : TForall τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iApply ("HKK" $! (TLamV _, TLamV _)); iFrame; iAlways. iClear "HKK".
iIntros (τi ? KK' j') "[Hv #HKK] /=".
iApply wp_tapp; iNext.
iMod (step_tlam with "[-]") as "Hz"; eauto.
iApply (`'IHHtyped (τi :: Δ)); iFrame "#"; iFrame.
by rewrite interp_env_ren.
Qed.
Lemma bin_log_related_tapp Γ e e' τ τ'
(IHHtyped : Γ ⊨ e ≤log≤ e' : TForall τ) :
Γ ⊨ TApp e ≤log≤ TApp e' : τ.[τ'/].
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (TAppCtx :: KK.1, TAppCtx :: KK.2)
vv j1 "[Hj1 #Hvv]". fold interp; simpl.
iSpecialize ("Hvv" $! (interp τ' Δ) with "[#]"); [iPureIntro; apply _|].
iApply "Hvv"; iFrame; eauto.
iAlways. iIntros (??) "?". iApply "HKK".
erewrite <- (interp_subst _ _ _ _); eauto.
Qed.
Lemma bin_log_related_fold Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ.[(TRec τ)/]) :
Γ ⊨ Fold e ≤log≤ Fold e' : TRec τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (FoldCtx :: KK.1, FoldCtx :: KK.2)
vv j1 "[Hj1 #Hvv]".
destruct vv as [v1 v2].
iApply ("HKK" $! (FoldV _, FoldV _)); iFrame; simpl.
rewrite -> (fixpoint_unfold (interp_rec1 _ _) _).
rewrite /= -(interp_subst _ _ _ _).
iAlways; iExists (_, _); eauto.
Qed.
Lemma bin_log_related_unfold Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : TRec τ) :
Γ ⊨ Unfold e ≤log≤ Unfold e' : τ.[(TRec τ)/].
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (UnfoldCtx :: KK.1, UnfoldCtx :: KK.2)
vv j1 "[Hj1 #Hvv]".
destruct vv as [v1 v2].
rewrite /= (fixpoint_unfold (interp_rec1 _ _) _) /=.
change (fixpoint _) with (interp (TRec τ) Δ).
iDestruct "Hvv" as ([w w']) "#[% Hiz]"; simplify_eq/=.
iMod (step_Fold with "[-]") as "Hz"; eauto.
iApply wp_unfold; cbn; auto.
iNext.
iApply ("HKK" $! (w, w')); iFrame.
by rewrite <- (interp_subst _ _ _ _).
Qed.
Lemma bin_log_related_fork Γ e e'
(IHHtyped : Γ ⊨ e ≤log≤ e' : TUnit) :
Γ ⊨ Fork e ≤log≤ Fork e' : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iMod (step_fork with "[-]") as (j') "[Hj Hj']"; eauto.
iApply wp_fork; iSplitL "Hj"; iNext.
- iApply ("HKK" $! (UnitV, UnitV)); eauto.
- iApply wp_wand_l; iSplitR; [|iApply (`'IHHtyped _ _ _ _ ([], []))]; eauto.
iFrame; iFrame "#".
iAlways; iIntros ([]?) "[? [% %]]"; simpl in *; subst; simpl.
iApply wp_value; eauto. iExists UnitV; eauto.
Qed.
Lemma bin_log_related_alloc Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : τ) :
Γ ⊨ Alloc e ≤log≤ Alloc e' : Tref τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (AllocCtx :: KK.1, AllocCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iMod (step_alloc with "[-]") as (l') "[Hj Hl']"; eauto.
iApply wp_alloc; eauto.
iNext. iIntros (l) "Hl".
iMod (inv_alloc (logN .@ (l,l')) _ (∃ w : val * val,
l ↦ᵢ w.1 ∗ l' ↦ₛ w.2 ∗ interp τ Δ w)%I with "[Hl Hl']") as "HN"; eauto.
{ iNext. iExists vv; iFrame; iFrame "#". }
iApply ("HKK" $! (LocV _, LocV _)); iFrame.
iExists (_, _); eauto.
Qed.
Lemma bin_log_related_load Γ e e' τ
(IHHtyped : Γ ⊨ e ≤log≤ e' : (Tref τ)) :
Γ ⊨ Load e ≤log≤ Load e' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped (LoadCtx :: KK.1, LoadCtx :: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([w w']) "[Hw1 [>Hw2 #Hw]]" "Hclose"; simpl.
iModIntro.
iApply (wp_load' with "[-]"); iFrame.
iNext. iIntros "Hw1".
iMod (step_load with "[Hj1 Hw2]") as "[Hj1 Hw2]";
[solve_ndisj|by iFrame|].
iApply wp_value; eauto.
iMod ("Hclose" with "[Hw1 Hw2]") as "_".
{ iNext. iExists (w,w'); by iFrame. }
iModIntro. iApply ("HKK" $! (w, w')); eauto.
Qed.
Lemma bin_log_related_store Γ e1 e2 e1' e2' τ
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : (Tref τ))
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ) :
Γ ⊨ Store e1 e2 ≤log≤ Store e1' e2' : TUnit.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (StoreLCtx _ :: KK.1, StoreLCtx _:: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (StoreRCtx _ :: KK.1, StoreRCtx _:: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([v v']) "[Hv1 [>Hv2 #Hv]]" "Hclose".
iModIntro.
iApply (wp_store' with "[-]"); auto using to_of_val.
iFrame. iNext. iIntros "Hv1".
iMod (step_store with "[$Hs Hj2 Hv2]") as "[Hj2 Hv2]"; eauto;
[solve_ndisj | by iFrame|].
iApply wp_value; eauto.
iMod ("Hclose" with "[Hv1 Hv2]") as "_".
{ iNext; iExists _; simpl; iFrame; iFrame "#". }
iModIntro. iApply ("HKK" $! (UnitV, UnitV)); eauto.
Qed.
Lemma bin_log_related_CAS Γ e1 e2 e3 e1' e2' e3' τ
(HEqτ : EqType τ)
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : Tref τ)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : τ)
(IHHtyped3 : Γ ⊨ e3 ≤log≤ e3' : τ) :
Γ ⊨ CAS e1 e2 e3 ≤log≤ CAS e1' e2' e3' : TBool.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (CasLCtx _ _ :: KK.1, CasLCtx _ _:: KK.2)
vv j1 "[Hj1 #Hvv]"; simpl.
smart_bind IHHtyped2 (CasMCtx _ _ :: KK.1, CasMCtx _ _:: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
smart_bind IHHtyped3 (CasRCtx _ _ :: KK.1, CasRCtx _ _:: KK.2)
uu j3 "[Hj3 #Huu]"; simpl.
destruct ww as [w w']; destruct uu as [u u'].
iDestruct "Hvv" as ([l l']) "[% Hll]"; simplify_eq/=.
iApply wp_atomic_under_ectx; eauto.
iInv (logN .@ (l,l')) as ([v v']) "[Hv1 [>Hv2 #Hv]]" "Hclose".
iModIntro.
iPoseProof ("Hv") as "Hv'".
rewrite {2}[⟦ τ ⟧ Δ (v, v')]interp_EqType_agree; trivial.
iMod "Hv'" as %Hv'; subst.
destruct (decide (v' = w)) as [|Hneq]; subst.
- iAssert (▷ ⌜w' = w⌝)%I as ">%".
{ rewrite ?interp_EqType_agree; trivial. by iSimplifyEq. }
simpl. iApply (wp_cas_suc' with "[-]"); eauto using to_of_val.
iFrame. iNext. iIntros "Hv1".
iMod (step_cas_suc
with "[Hj3 Hv2]") as "[Hj3 Hv2]"; simpl; eauto; first solve_ndisj.
iFrame. iFrame "Hs".
iApply wp_value; eauto.
iMod ("Hclose" with "[Hv1 Hv2]").
{ iNext; iExists (_, _); by iFrame. }
iModIntro. iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
- iAssert (▷ ⌜v' ≠ w'⌝)%I as ">%".
{ rewrite ?interp_EqType_agree; trivial. iSimplifyEq. auto. }
simpl. iApply (wp_cas_fail' with "[-]"); eauto.
iFrame. iNext. iIntros "Hv1".
iApply wp_value; eauto.
iMod (step_cas_fail
with "[$Hs Hj3 Hv2]") as "[Hj3 Hv2]"; simpl; eauto; first solve_ndisj.
iFrame.
iMod ("Hclose" with "[Hv1 Hv2]").
{ iNext; iExists (_, _); by iFrame. }
iModIntro. iApply ("HKK" $! (#♭v _, #♭v _)); eauto.
Qed.
Lemma bin_log_related_callcc Γ e e' τ
(Hclosed : ∀ f, e.[upn (S (length Γ)) f] = e)
(Hclosed' : ∀ f, e'.[upn (S (length Γ)) f] = e')
(IHHtyped : TCont τ :: Γ ⊨ e ≤log≤ e' : τ) :
Γ ⊨ Callcc e ≤log≤ Callcc e' : τ.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
iDestruct (interp_env_length with "HΓ") as %?.
iMod (step_callcc with "[Hj]") as "Hj"; eauto.
iApply wp_callcc; iNext.
asimpl.
change (Cont (KK.2)) with (of_val (ContV (KK.2))).
change (Cont (KK.1)) with (of_val (ContV (KK.1))).
erewrite !n_closed_subst_head_simpl by (rewrite ?fmap_length; eauto).
iApply (`'IHHtyped _ ((ContV (KK.1), ContV (KK.2)) :: vvs)).
iFrame; iFrame "#".
iApply interp_env_cons; iFrame "#"; auto.
simpl; iExists _, _; eauto.
Qed.
Lemma bin_log_related_throw Γ e1 e2 e1' e2' τ τ'
(IHHtyped1 : Γ ⊨ e1 ≤log≤ e1' : τ)
(IHHtyped2 : Γ ⊨ e2 ≤log≤ e2' : TCont τ) :
Γ ⊨ Throw e1 e2 ≤log≤ Throw e1' e2' : τ'.
Proof.
iIntros (Δ vvs ρ ?) "#(Hs & HΓ)"; iIntros (KK j) "[Hj #HKK] /=".
smart_bind IHHtyped1 (ThrowLCtx _ :: KK.1, ThrowLCtx _ :: KK.2)
vv j1 "[Hj1 #Hvv]". simpl.
smart_bind IHHtyped2 (ThrowRCtx _ :: KK.1, ThrowRCtx _ :: KK.2)
ww j2 "[Hj2 #Hww]"; simpl.
destruct ww as [w1 w2].
iDestruct "Hww" as (K K') "[% #HKK2]"; simplify_eq; simpl.
iMod (step_throw with "[Hj2]") as "Hj2"; eauto.
iApply wp_throw; eauto.
iNext.
iApply "HKK2"; eauto.
Qed.
Theorem binary_fundamental Γ e τ :
Γ ⊢ₜ e : τ → Γ ⊨ e ≤log≤ e : τ.
Proof.
induction 1.
- by apply bin_log_related_var.
- by apply bin_log_related_unit.
- by apply bin_log_related_nat.
- by apply bin_log_related_bool.
- apply bin_log_related_nat_binop; eauto.
- apply bin_log_related_pair; eauto.
- eapply bin_log_related_fst; eauto.
- eapply bin_log_related_snd; eauto.
- eapply bin_log_related_injl; eauto.
- eapply bin_log_related_injr; eauto.
- eapply bin_log_related_case; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_if; eauto.
- eapply bin_log_related_rec; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_Lam; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_LetIn; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_Seq; eauto.
- eapply (bin_log_related_app _ _ _ _ _ τ1); auto.
- eapply bin_log_related_tlam; eauto with typeclass_instances.
- eapply bin_log_related_tapp; eauto.
- eapply bin_log_related_fold; eauto.
- eapply bin_log_related_unfold; eauto.
- eapply bin_log_related_fork; eauto.
- eapply bin_log_related_alloc; eauto.
- eapply bin_log_related_load; eauto.
- eapply (bin_log_related_store _ _ _ _ _ τ) ; eauto.
- eapply bin_log_related_CAS; eauto.
- eapply bin_log_related_callcc; eauto;
match goal with H : _ |- _ => eapply (typed_n_closed _ _ _ H) end.
- eapply bin_log_related_throw; eauto.
Qed.
End fundamental.