RunST.rules
From iris.base_logic Require Export invariants big_op.
From iris.program_logic Require Import ectx_lifting.
From iris.algebra Require Import excl gmap gset cmra_big_op.
From iris.base_logic Require Import auth own.
From RunST Require Import lang ST_Lang_reduction.
From RunST Require Import IC IC_lifting ICTriple.
From iris.proofmode Require Import tactics.
Import uPred.
Section Two_ICs.
Context `{ICState ST_lang, ICG ST_lang Σ}.
Lemma ic_runst {γ} e (Φ : val → nat → iProp Σ) :
IC e {{v; n,
IC (of_val v : eff_expr)
{{ ev; m, IC RunST (of_eff_val ev) {{w; k, Φ w (n + m + k)}}[γ] }}[γ]
}}[γ]
⊢ IC RunST e {{Φ}}[γ].
Proof.
iIntros "Hic".
iApply (@ic_bind _ _ _ _ (fill_item RunSTCtx)); simpl.
iApply ic_mono; last eauto; simpl; iIntros (v n) "Hic".
rewrite ic_eq /ic_def.
iIntros (σ1 σ2 u n') "[Hr Ho]"; iDestruct "Hr" as %Hr.
destruct (nsteps_bind (fill_item RunSTCtx) Hr) as
(k & σ3 & w & Hkn & Hstp1 & Hstp2).
apply nsteps_val in Hstp1; destruct Hstp1 as [? [Hvls ?]]; subst; simpl in *.
replace (n' - 0) with n' in Hstp2 by omega.
destruct (red_runST _ _ _ _ _ Hstp2) as
(k' & z & σ'' & Hle & Hstp21 & Hstp22).
iSpecialize ("Hic" $! σ3 σ'' z k' with "[Ho]");
first by rewrite -Hvls; iFrame.
iMod "Hic" as "[Hic Ho]".
iSpecialize ("Hic" $! σ'' σ2 u (n' - k') with "[Ho]"); first by iFrame.
iMod "Hic".
iModIntro.
by replace (n + k' + (n' - k')) with (n + n') by omega.
Qed.
Lemma ic_runst_to_bind {γ} v w (Φ : val → nat → iProp Σ) :
IC RunST (of_val v) {{u; n,
IC RunST (App (of_val w) (of_val u)) {{ u'; m, Φ u' (n + m)}}[γ]
}}[γ]
⊢ IC RunST (Bind (of_val v) (of_val w)) {{Φ}}[γ].
Proof.
iIntros "Hic". rewrite ic_eq /ic_def.
iIntros (σ1 σ2 u n) "[Hred Ho]"; iDestruct "Hred" as %Hred.
destruct (red_RunST_Bind _ _ _ _ _ _ Hred) as (k & u' & σ' & Hle & Hrd1 & Hrd2).
iMod ("Hic" $! σ1 σ' u' k with "[$Ho]") as "[Hic Ho]"; auto.
iMod ("Hic" $! σ' σ2 u (n - k) with "[$Ho]") as "[Hic Ho]"; auto.
replace (k + (n - k)) with n by omega. by iModIntro; iFrame.
Qed.
End Two_ICs.
Definition heapUR : ucmraT := gmapUR loc (exclR valC).
Definition to_heap : state → heapUR := fmap (λ v, Excl v).
From iris.program_logic Require Import ectx_lifting.
From iris.algebra Require Import excl gmap gset cmra_big_op.
From iris.base_logic Require Import auth own.
From RunST Require Import lang ST_Lang_reduction.
From RunST Require Import IC IC_lifting ICTriple.
From iris.proofmode Require Import tactics.
Import uPred.
Section Two_ICs.
Context `{ICState ST_lang, ICG ST_lang Σ}.
Lemma ic_runst {γ} e (Φ : val → nat → iProp Σ) :
IC e {{v; n,
IC (of_val v : eff_expr)
{{ ev; m, IC RunST (of_eff_val ev) {{w; k, Φ w (n + m + k)}}[γ] }}[γ]
}}[γ]
⊢ IC RunST e {{Φ}}[γ].
Proof.
iIntros "Hic".
iApply (@ic_bind _ _ _ _ (fill_item RunSTCtx)); simpl.
iApply ic_mono; last eauto; simpl; iIntros (v n) "Hic".
rewrite ic_eq /ic_def.
iIntros (σ1 σ2 u n') "[Hr Ho]"; iDestruct "Hr" as %Hr.
destruct (nsteps_bind (fill_item RunSTCtx) Hr) as
(k & σ3 & w & Hkn & Hstp1 & Hstp2).
apply nsteps_val in Hstp1; destruct Hstp1 as [? [Hvls ?]]; subst; simpl in *.
replace (n' - 0) with n' in Hstp2 by omega.
destruct (red_runST _ _ _ _ _ Hstp2) as
(k' & z & σ'' & Hle & Hstp21 & Hstp22).
iSpecialize ("Hic" $! σ3 σ'' z k' with "[Ho]");
first by rewrite -Hvls; iFrame.
iMod "Hic" as "[Hic Ho]".
iSpecialize ("Hic" $! σ'' σ2 u (n' - k') with "[Ho]"); first by iFrame.
iMod "Hic".
iModIntro.
by replace (n + k' + (n' - k')) with (n + n') by omega.
Qed.
Lemma ic_runst_to_bind {γ} v w (Φ : val → nat → iProp Σ) :
IC RunST (of_val v) {{u; n,
IC RunST (App (of_val w) (of_val u)) {{ u'; m, Φ u' (n + m)}}[γ]
}}[γ]
⊢ IC RunST (Bind (of_val v) (of_val w)) {{Φ}}[γ].
Proof.
iIntros "Hic". rewrite ic_eq /ic_def.
iIntros (σ1 σ2 u n) "[Hred Ho]"; iDestruct "Hred" as %Hred.
destruct (red_RunST_Bind _ _ _ _ _ _ Hred) as (k & u' & σ' & Hle & Hrd1 & Hrd2).
iMod ("Hic" $! σ1 σ' u' k with "[$Ho]") as "[Hic Ho]"; auto.
iMod ("Hic" $! σ' σ2 u (n - k) with "[$Ho]") as "[Hic Ho]"; auto.
replace (k + (n - k)) with n by omega. by iModIntro; iFrame.
Qed.
End Two_ICs.
Definition heapUR : ucmraT := gmapUR loc (exclR valC).
Definition to_heap : state → heapUR := fmap (λ v, Excl v).
The CMRA we need.
Class heapG Σ := HeapG {
heap_invG :> invG Σ;
heap_inG :> authG Σ heapUR;
}.
Global Instance ICState_state : ICState ST_lang :=
{| IC_stateR := heapUR; IC_to_stateR := to_heap; |}.
Global Instance heap_G_ICG `{heapG Σ} : ICG ST_lang Σ := {|
ICstate_inG := heap_inG
|}.
Section definitions.
Context `{heapG Σ}.
Definition heap_mapsto_def γ (l : loc) (v: val) : iProp Σ :=
auth_own γ ({[ l := Excl v ]}).
Definition heap_mapsto_aux : { x | x = @heap_mapsto_def }. by eexists. Qed.
Definition heap_mapsto := proj1_sig heap_mapsto_aux.
Definition heap_mapsto_eq : @heap_mapsto = @heap_mapsto_def :=
proj2_sig heap_mapsto_aux.
End definitions.
Typeclasses Opaque heap_mapsto.
Notation "l ↦{ γ } v" := (heap_mapsto γ l v)
(at level 20, γ at level 50, format "l ↦{ γ } v") : uPred_scope.
Ltac inv_head_step :=
repeat match goal with
| _ => progress simplify_map_eq/= (* simplify memory stuff *)
| H : to_val _ = Some _ |- _ => apply of_to_val in H
| H : _ = ectx_language.fill ?K ?e |- _ => destruct K; simpl in *; subst
| H : _ = fill ?K ?e |- _ => destruct K; simpl in *; subst
| H : _ = eff_fill_item ?e (fill _ _) |- _ => destruct e; simpl in *; try congruence
| H : _ = fill_item ?e (fill _ _) |- _ => destruct e; simpl in *; try congruence
| H : head_step ?e _ _ _ _ |- _ =>
try (is_var e; fail 1); (* inversion yields many goals if e is a variable
and can thus better be avoided. *)
inversion H; subst; clear H
end.
Definition dec_atomic (e : eff_expr) :=
match e with
| Alloc e => bool_decide (is_Some (to_val e))
| Read e => bool_decide (is_Some (to_val e))
| Write e1 e2 => bool_decide (is_Some (to_val e1) ∧ is_Some (to_val e2))
| _ => false
end.
Lemma det_atomic_correct (e : eff_expr) : dec_atomic e → language.atomic e.
Proof.
intros He. apply ectx_language_atomic.
- intros σ e' σ' ef.
destruct e; simpl; try done; repeat (case_match; try done);
inversion 1; eauto;
match goal with
|- context[(Return (of_val ?v))] =>
change (Return (of_val v)) with (of_eff_val (ReturnEV v))
end; rewrite ?to_of_eff_val; eauto.
- intros [|Ki K] e' Hfill Hnotval; [done|exfalso].
apply (fill_not_val K), eq_None_not_Some in Hnotval. apply Hnotval. simpl.
revert He. destruct e; simpl; try done; repeat (case_match; try done);
rewrite ?bool_decide_spec;
destruct Ki; inversion Hfill; subst; clear Hfill;
try naive_solver eauto using to_val_is_Some.
Qed.
Section heap.
Context {Σ : gFunctors}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → nat → iProp Σ.
Implicit Types ψ : eff_val → nat → iProp Σ.
Implicit Types σ : state.
Implicit Types h g : heapUR.
heap_invG :> invG Σ;
heap_inG :> authG Σ heapUR;
}.
Global Instance ICState_state : ICState ST_lang :=
{| IC_stateR := heapUR; IC_to_stateR := to_heap; |}.
Global Instance heap_G_ICG `{heapG Σ} : ICG ST_lang Σ := {|
ICstate_inG := heap_inG
|}.
Section definitions.
Context `{heapG Σ}.
Definition heap_mapsto_def γ (l : loc) (v: val) : iProp Σ :=
auth_own γ ({[ l := Excl v ]}).
Definition heap_mapsto_aux : { x | x = @heap_mapsto_def }. by eexists. Qed.
Definition heap_mapsto := proj1_sig heap_mapsto_aux.
Definition heap_mapsto_eq : @heap_mapsto = @heap_mapsto_def :=
proj2_sig heap_mapsto_aux.
End definitions.
Typeclasses Opaque heap_mapsto.
Notation "l ↦{ γ } v" := (heap_mapsto γ l v)
(at level 20, γ at level 50, format "l ↦{ γ } v") : uPred_scope.
Ltac inv_head_step :=
repeat match goal with
| _ => progress simplify_map_eq/= (* simplify memory stuff *)
| H : to_val _ = Some _ |- _ => apply of_to_val in H
| H : _ = ectx_language.fill ?K ?e |- _ => destruct K; simpl in *; subst
| H : _ = fill ?K ?e |- _ => destruct K; simpl in *; subst
| H : _ = eff_fill_item ?e (fill _ _) |- _ => destruct e; simpl in *; try congruence
| H : _ = fill_item ?e (fill _ _) |- _ => destruct e; simpl in *; try congruence
| H : head_step ?e _ _ _ _ |- _ =>
try (is_var e; fail 1); (* inversion yields many goals if e is a variable
and can thus better be avoided. *)
inversion H; subst; clear H
end.
Definition dec_atomic (e : eff_expr) :=
match e with
| Alloc e => bool_decide (is_Some (to_val e))
| Read e => bool_decide (is_Some (to_val e))
| Write e1 e2 => bool_decide (is_Some (to_val e1) ∧ is_Some (to_val e2))
| _ => false
end.
Lemma det_atomic_correct (e : eff_expr) : dec_atomic e → language.atomic e.
Proof.
intros He. apply ectx_language_atomic.
- intros σ e' σ' ef.
destruct e; simpl; try done; repeat (case_match; try done);
inversion 1; eauto;
match goal with
|- context[(Return (of_val ?v))] =>
change (Return (of_val v)) with (of_eff_val (ReturnEV v))
end; rewrite ?to_of_eff_val; eauto.
- intros [|Ki K] e' Hfill Hnotval; [done|exfalso].
apply (fill_not_val K), eq_None_not_Some in Hnotval. apply Hnotval. simpl.
revert He. destruct e; simpl; try done; repeat (case_match; try done);
rewrite ?bool_decide_spec;
destruct Ki; inversion Hfill; subst; clear Hfill;
try naive_solver eauto using to_val_is_Some.
Qed.
Section heap.
Context {Σ : gFunctors}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → nat → iProp Σ.
Implicit Types ψ : eff_val → nat → iProp Σ.
Implicit Types σ : state.
Implicit Types h g : heapUR.
Conversion to heaps and back
Lemma to_heap_valid σ : ✓ to_heap σ.
Proof. intros l. rewrite lookup_fmap. by case (σ !! l). Qed.
Lemma to_heap_fullstate_valid σ : ✓ (● (to_heap σ) ⋅ ◯ (to_heap σ)).
Proof.
apply auth_valid_discrete_2; split; first done.
apply to_heap_valid.
Qed.
Lemma lookup_to_heap_None σ l : σ !! l = None → to_heap σ !! l = None.
Proof. by rewrite /to_heap lookup_fmap=> ->. Qed.
Lemma lookup_to_heap_Some σ l v : σ !! l = Some v → to_heap σ !! l = Excl' v.
Proof. by rewrite /to_heap lookup_fmap=> ->. Qed.
Lemma heap_singleton_included σ l v :
{[l := Excl v]} ≼ to_heap σ → σ !! l = Some v.
Proof.
rewrite singleton_included=> -[av [/leibniz_equiv_iff Hl Hqv]].
move: Hl. rewrite /to_heap lookup_fmap.
case: (σ !! l); simpl; last congruence; move => ? Heq; inversion Heq; subst.
by unshelve erewrite <- (f_equal _ (leibniz_equiv v _ (Excl_included _ _ _))).
Qed.
Lemma heap_singleton_included' σ l v :
{[l := Excl v]} ≼ to_heap σ → to_heap σ !! l = Some (Excl v).
Proof.
intros Hl%heap_singleton_included. by rewrite /to_heap lookup_fmap Hl.
Qed.
Lemma to_heap_insert l v σ :
to_heap (<[l:=v]> σ) = <[l:= Excl v]> (to_heap σ).
Proof. by rewrite /to_heap fmap_insert. Qed.
Context `{heapG Σ}.
Lemma heap_included γ σ1 σ2:
ownP_full γ σ1 ★ auth_own γ (to_heap σ2) ⊢ ■ (to_heap σ2 ≼ to_heap σ1 ∧ ✓ to_heap σ1).
Proof.
iIntros "(H1 & H2)". rewrite /ownP_full /auth_own.
iDestruct (own_valid_2 with "[$H1 $H2]") as %Hvl.
iFrame. apply auth_valid_discrete_2 in Hvl.
destruct Hvl as [Hvl1 Hvl2]; simpl in *; auto.
Qed.
Lemma heap_catch_up σ1 σ2 γ :
✓ to_heap σ1 →
(to_heap σ2 ≼ to_heap σ1) →
ownP_full γ σ2 ==★ ownP_full γ σ1.
Proof.
iIntros (Hvl Hinc) "H". rewrite /ownP_full.
iMod (own_update _ _ (● to_heap σ1) with "[$H]") as "H"; auto.
intros n z Hz.
destruct z as [[z1 z2]|]; simpl in *; last first.
{ constructor; simpl; auto. apply ucmra_unit_leastN. }
destruct z1 as [[z1|]|].
{ revert Hz; rewrite ?(auth_both_op z1 z2) => Hz; inversion Hz. }
{ inversion Hz. }
revert Hz; rewrite /validN /cmra_validN /= left_id. intros [Hz1 Hz2].
split; auto. etrans; first apply Hz1. by apply cmra_included_includedN.
Qed.
Lemma allocate_full_state σ :
True ==★ ∃ γ, ownP_full γ σ ★ auth_own γ (to_heap σ).
Proof.
iIntros.
iMod (own_alloc) as (γ) "Ho"; first apply (to_heap_fullstate_valid σ).
iModIntro; iExists γ. by rewrite own_op /ownP_full /auth_own.
Qed.
Lemma mapsto_alloc γ σ l v :
σ !! l = None →
ownP_full γ σ ==★ ownP_full γ (<[l := v]>σ) ★ l ↦{γ} v.
Proof.
intros Hl. iIntros "Ho".
rewrite /ownP_full heap_mapsto_eq /heap_mapsto_def /IC_to_stateR /=.
iMod (own_update with "Ho") as "H".
- apply auth_update_alloc.
eapply (alloc_singleton_local_update _ _ (Excl v)); last done.
apply lookup_to_heap_None; eauto.
- rewrite own_op; rewrite /auth_own. by rewrite to_heap_insert.
Qed.
Lemma mapsto_allocated γ σ l v :
ownP_full γ σ ★ l ↦{γ} v ⊢ ■ (σ !! l = Some v).
Proof.
iIntros "[H1 H2]".
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
iDestruct (own_valid_2 with "[$H1 $H2]") as %Hv%auth_valid_discrete_2.
destruct Hv as [Hv _]. by apply heap_singleton_included in Hv.
Qed.
Lemma mapsto_update γ σ l v w :
ownP_full γ σ ★ l ↦{γ} v ==★ ownP_full γ (<[l := w]>σ) ★ l ↦{γ} w.
Proof.
iIntros "[H1 H2]".
iDestruct (mapsto_allocated with "[$H1 $H2]") as %Hex.
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
iMod (own_update_2 with "[$H1 $H2]") as "H".
- apply auth_update.
eapply (singleton_local_update _ _ (Excl v) (Excl v) (Excl w) (Excl w)).
+ by apply lookup_to_heap_Some.
+ by apply exclusive_local_update.
- by rewrite own_op to_heap_insert.
Qed.
Lemma mapsto_no_dup γ l v w : l ↦{γ} v ★ l ↦{γ} w ⊢ False.
Proof.
iIntros "H".
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
set (W:= auth_own_op); unfold auth_own in W; rewrite -W; clear W.
rewrite op_singleton.
iDestruct (own_valid with "H") as %Hv.
specialize (Hv l); revert Hv; rewrite /= lookup_singleton; inversion 1.
Qed.
Proof. intros l. rewrite lookup_fmap. by case (σ !! l). Qed.
Lemma to_heap_fullstate_valid σ : ✓ (● (to_heap σ) ⋅ ◯ (to_heap σ)).
Proof.
apply auth_valid_discrete_2; split; first done.
apply to_heap_valid.
Qed.
Lemma lookup_to_heap_None σ l : σ !! l = None → to_heap σ !! l = None.
Proof. by rewrite /to_heap lookup_fmap=> ->. Qed.
Lemma lookup_to_heap_Some σ l v : σ !! l = Some v → to_heap σ !! l = Excl' v.
Proof. by rewrite /to_heap lookup_fmap=> ->. Qed.
Lemma heap_singleton_included σ l v :
{[l := Excl v]} ≼ to_heap σ → σ !! l = Some v.
Proof.
rewrite singleton_included=> -[av [/leibniz_equiv_iff Hl Hqv]].
move: Hl. rewrite /to_heap lookup_fmap.
case: (σ !! l); simpl; last congruence; move => ? Heq; inversion Heq; subst.
by unshelve erewrite <- (f_equal _ (leibniz_equiv v _ (Excl_included _ _ _))).
Qed.
Lemma heap_singleton_included' σ l v :
{[l := Excl v]} ≼ to_heap σ → to_heap σ !! l = Some (Excl v).
Proof.
intros Hl%heap_singleton_included. by rewrite /to_heap lookup_fmap Hl.
Qed.
Lemma to_heap_insert l v σ :
to_heap (<[l:=v]> σ) = <[l:= Excl v]> (to_heap σ).
Proof. by rewrite /to_heap fmap_insert. Qed.
Context `{heapG Σ}.
Lemma heap_included γ σ1 σ2:
ownP_full γ σ1 ★ auth_own γ (to_heap σ2) ⊢ ■ (to_heap σ2 ≼ to_heap σ1 ∧ ✓ to_heap σ1).
Proof.
iIntros "(H1 & H2)". rewrite /ownP_full /auth_own.
iDestruct (own_valid_2 with "[$H1 $H2]") as %Hvl.
iFrame. apply auth_valid_discrete_2 in Hvl.
destruct Hvl as [Hvl1 Hvl2]; simpl in *; auto.
Qed.
Lemma heap_catch_up σ1 σ2 γ :
✓ to_heap σ1 →
(to_heap σ2 ≼ to_heap σ1) →
ownP_full γ σ2 ==★ ownP_full γ σ1.
Proof.
iIntros (Hvl Hinc) "H". rewrite /ownP_full.
iMod (own_update _ _ (● to_heap σ1) with "[$H]") as "H"; auto.
intros n z Hz.
destruct z as [[z1 z2]|]; simpl in *; last first.
{ constructor; simpl; auto. apply ucmra_unit_leastN. }
destruct z1 as [[z1|]|].
{ revert Hz; rewrite ?(auth_both_op z1 z2) => Hz; inversion Hz. }
{ inversion Hz. }
revert Hz; rewrite /validN /cmra_validN /= left_id. intros [Hz1 Hz2].
split; auto. etrans; first apply Hz1. by apply cmra_included_includedN.
Qed.
Lemma allocate_full_state σ :
True ==★ ∃ γ, ownP_full γ σ ★ auth_own γ (to_heap σ).
Proof.
iIntros.
iMod (own_alloc) as (γ) "Ho"; first apply (to_heap_fullstate_valid σ).
iModIntro; iExists γ. by rewrite own_op /ownP_full /auth_own.
Qed.
Lemma mapsto_alloc γ σ l v :
σ !! l = None →
ownP_full γ σ ==★ ownP_full γ (<[l := v]>σ) ★ l ↦{γ} v.
Proof.
intros Hl. iIntros "Ho".
rewrite /ownP_full heap_mapsto_eq /heap_mapsto_def /IC_to_stateR /=.
iMod (own_update with "Ho") as "H".
- apply auth_update_alloc.
eapply (alloc_singleton_local_update _ _ (Excl v)); last done.
apply lookup_to_heap_None; eauto.
- rewrite own_op; rewrite /auth_own. by rewrite to_heap_insert.
Qed.
Lemma mapsto_allocated γ σ l v :
ownP_full γ σ ★ l ↦{γ} v ⊢ ■ (σ !! l = Some v).
Proof.
iIntros "[H1 H2]".
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
iDestruct (own_valid_2 with "[$H1 $H2]") as %Hv%auth_valid_discrete_2.
destruct Hv as [Hv _]. by apply heap_singleton_included in Hv.
Qed.
Lemma mapsto_update γ σ l v w :
ownP_full γ σ ★ l ↦{γ} v ==★ ownP_full γ (<[l := w]>σ) ★ l ↦{γ} w.
Proof.
iIntros "[H1 H2]".
iDestruct (mapsto_allocated with "[$H1 $H2]") as %Hex.
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
iMod (own_update_2 with "[$H1 $H2]") as "H".
- apply auth_update.
eapply (singleton_local_update _ _ (Excl v) (Excl v) (Excl w) (Excl w)).
+ by apply lookup_to_heap_Some.
+ by apply exclusive_local_update.
- by rewrite own_op to_heap_insert.
Qed.
Lemma mapsto_no_dup γ l v w : l ↦{γ} v ★ l ↦{γ} w ⊢ False.
Proof.
iIntros "H".
rewrite /ownP_full heap_mapsto_eq
/heap_mapsto_def /IC_to_stateR /auth_own /=.
set (W:= auth_own_op); unfold auth_own in W; rewrite -W; clear W.
rewrite op_singleton.
iDestruct (own_valid with "H") as %Hv.
specialize (Hv l); revert Hv; rewrite /= lookup_singleton; inversion 1.
Qed.
General properties of mapsto
Global Instance heap_mapsto_timeless γ l v : TimelessP (l ↦{γ} v).
Proof. rewrite heap_mapsto_eq /heap_mapsto_def. apply _. Qed.
Global Instance heap_auth_timeless γ σ : TimelessP (ownP_full γ σ).
Proof. unfold ownP_full. apply _. Qed.
End heap.
Section lang_rules.
Context `{heapG Σ}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → nat → iProp Σ.
Implicit Types ψ : eff_val → nat → iProp Σ.
Implicit Types σ : state.
Implicit Types h g : heapUR.
Lemma ic_alloc γ E e v Ψ :
to_val e = Some v →
{{| ▷ ∀ l, l ↦{γ} v -★ Ψ (ReturnEV (LitV (LitLoc l))) 1 |}}
(Alloc e : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros (Heq) "!#HΨ".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. apply det_atomic_correct; rewrite /= to_of_val; by cbv. }
iIntros (σ1) "Ho"; iModIntro; iSplit; first iPureIntro.
{ eexists _, _. by apply alloc_fresh. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
match goal with
H1 : ?A = Some _, H2 : ?A = Some _ |- _ =>
rewrite H1 in H2; inversion H2; subst
end.
iMod (mapsto_alloc with "Ho") as "[Ho Hl]"; eauto.
iFrame; iModIntro. iNext.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_read γ E l v Ψ :
{{| ▷ l ↦{γ} v ★ ▷ (l ↦{γ} v -★ Ψ (ReturnEV v) 1) |}}
(Read (Lit (LitLoc l)) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#[Hl HΨ]".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. by apply det_atomic_correct; simpl. }
iIntros (σ1) "Ho"; iMod "Hl"; iModIntro.
iDestruct (mapsto_allocated with "[$Ho $Hl]") as %Hi.
iSplit.
{ iPureIntro. eexists _, _; econstructor; eauto. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
match goal with
H1 : ?A = Some _, H2 : ?A = Some _ |- _ =>
rewrite H1 in H2; inversion H2; subst
end.
iModIntro. iNext; iFrame.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
simpl in H'; rewrite to_of_val in H'; inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_write γ E l v w Ψ :
{{| ▷ l ↦{γ} v ★ ▷ (l ↦{γ} w -★ Ψ (ReturnEV (LitV LitUnit)) 1) |}}
(Write (Lit (LitLoc l)) (of_val w) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#[Hl HΨ]".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. by apply det_atomic_correct; rewrite /= to_of_val /=. }
iIntros (σ1) "Ho"; iMod "Hl"; iModIntro.
iDestruct (mapsto_allocated with "[$Ho $Hl]") as %Hi.
iSplit.
{ iPureIntro. eexists _, _; econstructor; rewrite ?to_of_val; eauto. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
iMod (mapsto_update with "[$Ho $Hl]") as "[Ho Hl]".
iModIntro; iNext; iFrame.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
inversion H'; subst
end.
match goal with
H' : to_val (of_val _ )= Some _ |- _ =>
rewrite to_of_val in H'; inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_eff_bind γ E v w Ψ :
{{| ▷ Ψ (AppCtxEV [] w v) 1 |}}
(Bind (Return (of_val v)) (of_val w) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#HΨ".
iApply (@ic_lift_pure_head_step _ _ _ _ eff_ectx_lang).
{ intros; eexists _, _; econstructor; by rewrite to_of_val. }
{ intros ? ? ? Hrd; by inversion Hrd; subst. }
iNext. iIntros (σ1 e2 σ2 Hrd). inversion Hrd; subst.
change (App (of_val w) (of_val v)) with
(@language.of_val eff_lang (AppCtxEV [] w v)).
by iApply (@ic_value' eff_lang).
Qed.
Lemma ic_return γ E v Φ :
{{| ▷ Φ v 1 |}} RunST (Return (of_val v)) @ E {{| Φ |}}[γ].
Proof.
iIntros "!#HΨ".
iApply (@ic_lift_pure_head_step _ _ _ _ ST_ectx_lang).
{ intros; eexists _, _; econstructor; by rewrite to_of_val. }
{ intros ? ? ? Hrd. inversion Hrd; subst; trivial.
match goal with
H' : prim_step (Return _) _ _ _ _ |- _ =>
apply language.val_stuck in H'; rewrite /= to_of_val /= in H';
inversion H'
end. }
iNext. iIntros (σ1 e2 σ2 Hrd).
inversion Hrd; subst.
{ match goal with
H' : prim_step (Return _) _ _ _ _ |- _ =>
apply language.val_stuck in H'; rewrite /= to_of_val /= in H';
inversion H'
end. }
by iApply ic_value'.
Qed.
Local Hint Extern 1 (nf_head_reducible _ _) =>
eexists _, _; econstructor; auto using to_of_val.
Lemma ic_rec γ E e1 e2 v Φ :
to_val e2 = Some v →
▷ IC e1.[(Rec e1), e2 /] @ E {{w; n, Φ w (S n) }}[γ]
⊢ IC App (Rec e1) e2 @ E {{ Φ }}[γ].
Proof.
intros <-%of_to_val.
rewrite -(ic_lift_pure_det_head_step (App _ _) e1.[Rec e1,of_val v/]);
eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_tlam γ E e Φ :
▷ IC e @ E {{w; n, Φ w (S n) }}[γ] ⊢ IC TApp (TLam e) @ E {{ Φ }}[γ].
Proof.
rewrite -(ic_lift_pure_det_head_step (TApp _) e); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_fold γ E e v Φ :
to_val e = Some v → ▷ (|={E}=> Φ v 1) ⊢ IC Unfold (Fold e) @ E {{ Φ }}[γ].
Proof.
intros <-%of_to_val.
rewrite -(ic_lift_pure_det_head_step (Unfold _) (of_val v))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_fst γ E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
▷ (|={E}=> Φ v1 1) ⊢ IC Fst (Pair e1 e2) @ E {{ Φ }}[γ].
Proof.
intros. rewrite -(ic_lift_pure_det_head_step (Fst _) e1)
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_snd γ E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
▷ (|={E}=> Φ v2 1) ⊢ IC Snd (Pair e1 e2) @ E {{ Φ }}[γ].
Proof.
intros. rewrite -(ic_lift_pure_det_head_step (Snd _) e2)
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_case_inl γ E e0 v0 e1 e2 Φ :
to_val e0 = Some v0 →
▷ IC e1.[e0/] @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC Case (InjL e0) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Case _ _ _) (e1.[e0/])); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_case_inr γ E e0 v0 e1 e2 Φ :
to_val e0 = Some v0 →
▷ IC e2.[e0/] @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC Case (InjR e0) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Case _ _ _) (e2.[e0/])); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_bin_op γ E op z1 z2 Φ :
▷ (|={E}=> Φ (bin_op_eval op z1 z2) 1)
⊢ IC BinOp op (Lit $ LitInt z1) (Lit $ LitInt z2) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step
(BinOp _ _ _) (of_val (bin_op_eval op z1 z2)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_if_true γ E e1 e2 Φ :
▷ IC e1 @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC If (Lit (LitBool true)) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (If _ _ _) (e1)); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_if_false γ E e1 e2 Φ :
▷ IC e2 @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC If (Lit (LitBool false)) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (If _ _ _) (e2)); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_compare_suc γ E l Φ :
▷ (|={E}=> Φ (LitV (LitBool true)) 1) ⊢
IC Compare (Lit (LitLoc l)) (Lit (LitLoc l)) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Compare _ _) (Lit (LitBool true)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_compare_fail γ E l1 l2 Φ :
l1 ≠ l2 →
▷ (|={E}=> Φ (LitV (LitBool false)) 1) ⊢
IC Compare (Lit (LitLoc l1)) (Lit (LitLoc l2)) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Compare _ _) (Lit (LitBool false)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
End lang_rules.
Proof. rewrite heap_mapsto_eq /heap_mapsto_def. apply _. Qed.
Global Instance heap_auth_timeless γ σ : TimelessP (ownP_full γ σ).
Proof. unfold ownP_full. apply _. Qed.
End heap.
Section lang_rules.
Context `{heapG Σ}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → nat → iProp Σ.
Implicit Types ψ : eff_val → nat → iProp Σ.
Implicit Types σ : state.
Implicit Types h g : heapUR.
Lemma ic_alloc γ E e v Ψ :
to_val e = Some v →
{{| ▷ ∀ l, l ↦{γ} v -★ Ψ (ReturnEV (LitV (LitLoc l))) 1 |}}
(Alloc e : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros (Heq) "!#HΨ".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. apply det_atomic_correct; rewrite /= to_of_val; by cbv. }
iIntros (σ1) "Ho"; iModIntro; iSplit; first iPureIntro.
{ eexists _, _. by apply alloc_fresh. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
match goal with
H1 : ?A = Some _, H2 : ?A = Some _ |- _ =>
rewrite H1 in H2; inversion H2; subst
end.
iMod (mapsto_alloc with "Ho") as "[Ho Hl]"; eauto.
iFrame; iModIntro. iNext.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_read γ E l v Ψ :
{{| ▷ l ↦{γ} v ★ ▷ (l ↦{γ} v -★ Ψ (ReturnEV v) 1) |}}
(Read (Lit (LitLoc l)) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#[Hl HΨ]".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. by apply det_atomic_correct; simpl. }
iIntros (σ1) "Ho"; iMod "Hl"; iModIntro.
iDestruct (mapsto_allocated with "[$Ho $Hl]") as %Hi.
iSplit.
{ iPureIntro. eexists _, _; econstructor; eauto. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
match goal with
H1 : ?A = Some _, H2 : ?A = Some _ |- _ =>
rewrite H1 in H2; inversion H2; subst
end.
iModIntro. iNext; iFrame.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
simpl in H'; rewrite to_of_val in H'; inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_write γ E l v w Ψ :
{{| ▷ l ↦{γ} v ★ ▷ (l ↦{γ} w -★ Ψ (ReturnEV (LitV LitUnit)) 1) |}}
(Write (Lit (LitLoc l)) (of_val w) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#[Hl HΨ]".
iApply (@ic_lift_atomic_head_step _ _ _ _ eff_ectx_lang).
{ inv_head_step. by apply det_atomic_correct; rewrite /= to_of_val /=. }
iIntros (σ1) "Ho"; iMod "Hl"; iModIntro.
iDestruct (mapsto_allocated with "[$Ho $Hl]") as %Hi.
iSplit.
{ iPureIntro. eexists _, _; econstructor; rewrite ?to_of_val; eauto. }
iIntros (v2 σ2 Hrd).
inversion Hrd; subst.
iMod (mapsto_update with "[$Ho $Hl]") as "[Ho Hl]".
iModIntro; iNext; iFrame.
match goal with
H : Return _ = of_eff_val v2 |- _ =>
let H' := fresh "H" in
pose proof (f_equal to_eff_val H) as H'; rewrite to_of_eff_val in H';
inversion H'; subst
end.
match goal with
H' : to_val (of_val _ )= Some _ |- _ =>
rewrite to_of_val in H'; inversion H'; subst
end.
by iApply "HΨ".
Qed.
Lemma ic_eff_bind γ E v w Ψ :
{{| ▷ Ψ (AppCtxEV [] w v) 1 |}}
(Bind (Return (of_val v)) (of_val w) : eff_expr) @ E {{| Ψ |}}[γ].
Proof.
iIntros "!#HΨ".
iApply (@ic_lift_pure_head_step _ _ _ _ eff_ectx_lang).
{ intros; eexists _, _; econstructor; by rewrite to_of_val. }
{ intros ? ? ? Hrd; by inversion Hrd; subst. }
iNext. iIntros (σ1 e2 σ2 Hrd). inversion Hrd; subst.
change (App (of_val w) (of_val v)) with
(@language.of_val eff_lang (AppCtxEV [] w v)).
by iApply (@ic_value' eff_lang).
Qed.
Lemma ic_return γ E v Φ :
{{| ▷ Φ v 1 |}} RunST (Return (of_val v)) @ E {{| Φ |}}[γ].
Proof.
iIntros "!#HΨ".
iApply (@ic_lift_pure_head_step _ _ _ _ ST_ectx_lang).
{ intros; eexists _, _; econstructor; by rewrite to_of_val. }
{ intros ? ? ? Hrd. inversion Hrd; subst; trivial.
match goal with
H' : prim_step (Return _) _ _ _ _ |- _ =>
apply language.val_stuck in H'; rewrite /= to_of_val /= in H';
inversion H'
end. }
iNext. iIntros (σ1 e2 σ2 Hrd).
inversion Hrd; subst.
{ match goal with
H' : prim_step (Return _) _ _ _ _ |- _ =>
apply language.val_stuck in H'; rewrite /= to_of_val /= in H';
inversion H'
end. }
by iApply ic_value'.
Qed.
Local Hint Extern 1 (nf_head_reducible _ _) =>
eexists _, _; econstructor; auto using to_of_val.
Lemma ic_rec γ E e1 e2 v Φ :
to_val e2 = Some v →
▷ IC e1.[(Rec e1), e2 /] @ E {{w; n, Φ w (S n) }}[γ]
⊢ IC App (Rec e1) e2 @ E {{ Φ }}[γ].
Proof.
intros <-%of_to_val.
rewrite -(ic_lift_pure_det_head_step (App _ _) e1.[Rec e1,of_val v/]);
eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_tlam γ E e Φ :
▷ IC e @ E {{w; n, Φ w (S n) }}[γ] ⊢ IC TApp (TLam e) @ E {{ Φ }}[γ].
Proof.
rewrite -(ic_lift_pure_det_head_step (TApp _) e); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_fold γ E e v Φ :
to_val e = Some v → ▷ (|={E}=> Φ v 1) ⊢ IC Unfold (Fold e) @ E {{ Φ }}[γ].
Proof.
intros <-%of_to_val.
rewrite -(ic_lift_pure_det_head_step (Unfold _) (of_val v))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_fst γ E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
▷ (|={E}=> Φ v1 1) ⊢ IC Fst (Pair e1 e2) @ E {{ Φ }}[γ].
Proof.
intros. rewrite -(ic_lift_pure_det_head_step (Fst _) e1)
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_snd γ E e1 v1 e2 v2 Φ :
to_val e1 = Some v1 → to_val e2 = Some v2 →
▷ (|={E}=> Φ v2 1) ⊢ IC Snd (Pair e1 e2) @ E {{ Φ }}[γ].
Proof.
intros. rewrite -(ic_lift_pure_det_head_step (Snd _) e2)
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_case_inl γ E e0 v0 e1 e2 Φ :
to_val e0 = Some v0 →
▷ IC e1.[e0/] @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC Case (InjL e0) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Case _ _ _) (e1.[e0/])); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_case_inr γ E e0 v0 e1 e2 Φ :
to_val e0 = Some v0 →
▷ IC e2.[e0/] @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC Case (InjR e0) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Case _ _ _) (e2.[e0/])); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_bin_op γ E op z1 z2 Φ :
▷ (|={E}=> Φ (bin_op_eval op z1 z2) 1)
⊢ IC BinOp op (Lit $ LitInt z1) (Lit $ LitInt z2) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step
(BinOp _ _ _) (of_val (bin_op_eval op z1 z2)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_if_true γ E e1 e2 Φ :
▷ IC e1 @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC If (Lit (LitBool true)) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (If _ _ _) (e1)); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_if_false γ E e1 e2 Φ :
▷ IC e2 @ E {{w; n, Φ w (S n)}}[γ]
⊢ IC If (Lit (LitBool false)) e1 e2 @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (If _ _ _) (e2)); eauto.
intros; inv_head_step; eauto.
Qed.
Lemma ic_compare_suc γ E l Φ :
▷ (|={E}=> Φ (LitV (LitBool true)) 1) ⊢
IC Compare (Lit (LitLoc l)) (Lit (LitLoc l)) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Compare _ _) (Lit (LitBool true)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
Lemma ic_compare_fail γ E l1 l2 Φ :
l1 ≠ l2 →
▷ (|={E}=> Φ (LitV (LitBool false)) 1) ⊢
IC Compare (Lit (LitLoc l1)) (Lit (LitLoc l2)) @ E {{ Φ }}[γ].
Proof.
intros.
rewrite -(ic_lift_pure_det_head_step (Compare _ _) (Lit (LitBool false)))
-?ic_value_fupd; eauto using to_of_val.
intros; inv_head_step; eauto.
Qed.
End lang_rules.