RunST.state_independence
From RunST Require Import IC lang rules typing regions logrel fundamental
reduction contextual_refinement IC_adequacy saved_pred ST_Lang_reduction
logrel_shared logrel_ctx_closure.
From iris.algebra Require Import gmap excl dec_agree.
From iris.base_logic Require Import big_op auth.
From iris.proofmode Require Import tactics.
Lemma special_reg_alloc E `{heapG Σ} `{PreLogRelG Σ} l l' ρ
τ :
True ={E}=★ ∃ _ : LogRelG Σ,
reg_inv ★ STRefREL ((@type_to_reg Σ ρ) [])
(⟦τ⟧ []) (LitV (LitLoc l), LitV (LitLoc l')).
Proof.
iIntros "".
iMod (own_alloc (● (∅ : regsUR))) as (γ) "Hregs"; first done.
set (LRG := PreLogRelG_LogRelG γ).
iExists (LRG).
rewrite reg_inv_eq /reg_inv_def /_reg_inv STRefREL_eq /STRefREL_def /=.
iMod (saved_pred_alloc_strong (⟦τ⟧ []) ∅) as (γΦ) "[_ #Hpred]".
iMod (alloc_empty_preds) as (γrel) "HPR".
iMod (Rel_name_alloc _ _ l l' γΦ with "HPR") as "[HPR #Hll]"; first done.
iAssert (Rel_for γrel l l' (⟦τ⟧ [])) as "#HRF";
first by rewrite Rel_for_eq /Rel_for_def /=; iExists _; iFrame "#".
iMod (alloc_empty_bij) as (γbij) "Hbij".
iMod (bij_alloc _ _ l l' with "Hbij") as "[Hbij Hlbij]"; try done.
rewrite union_empty_r_L.
iMod (PREDS_PREDS_include with "HPR") as "[HPR HPRi]".
iMod (own_update with "Hregs") as "Hregs".
{ apply auth_update_alloc.
by apply (alloc_singleton_local_update
_ 1%positive (DecAgree (γbij, γrel))). }
rewrite own_op. iDestruct "Hregs" as "[Hregs How]".
iMod (inv_alloc regIN E
(∃ Rs, REGS Rs ★
[★ map] n ↦ dR ∈ Rs, ∃ R, dR = DecAgree R ★
(∃ g, BIJ (fst R) g ★ PREDS_include (snd R) g))%I
with "[Hregs Hbij HPRi]") as "#Hinv".
- rewrite REGS_eq /REGS_def.
iNext; iExists _; iFrame.
rewrite big_sepM_insert; last done.
iSplitL; last by rewrite big_sepM_empty.
iExists (_, _); iSplit; first done; simpl.
iExists _; iFrame.
- iModIntro. iFrame "#".
iExists l, l', (γbij, γrel); simpl; iFrame; iFrame "#".
iSplit; trivial.
rewrite REG_eq /REG_def type_to_reg_eq /type_to_reg_def /=.
by destruct ρ.
Qed.
Lemma n_closed_subst_head_simpl_singleton e w :
(∀ f, e.[up f] = e) →
e.[of_val w/] = e.[env_subst [w]].
Proof.
intros H1.
rewrite /env_subst. eapply (n_closed_invariant 1); eauto=> /= -[|x] ? //=.
omega.
Qed.
Notation ICG_ST Σ := (@ICG' state ICState_state Σ).
Global Instance ICG_heapG `{H: ICG_ST Σ} : heapG Σ :=
{| heap_invG := IC_invG; heap_inG := ICstate_inG |}.
Lemma basic_state_independence Σ `{ICPreG ST_lang Σ} `{PreLogRelG Σ}
e σ1 e' σ1' ρ τ1 τ v1 σ2 l l' :
[TSTref ρ τ1] ⊢ₜ e : τ → [TSTref ρ τ1] ⊢ₜ e' : τ →
rtc pstep (e.[Lit $ LitLoc l/], σ1) (of_val v1, σ2) →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, [TSTref ρ τ1] ⊨ e ≤log≤ e' : τ) →
(∃ σ2' v2', rtc pstep (e'.[Lit $ LitLoc l'/], σ1') (of_val v2', σ2')
∧ (∀ i w, σ1' !! i = Some w → σ2' !! i = Some w)).
Proof.
intros Hte Hte' Hsteps Hlog.
apply rtc_nsteps in Hsteps. destruct Hsteps as [n Hsteps].
cut (∀ v n σ', nsteps pstep n (e.[Lit $ LitLoc l/], σ1) (of_val v, σ') →
((λ _ _, ∃ σ2' v2',
rtc pstep (e'.[Lit $ LitLoc l'/], σ1') (of_val v2', σ2')
∧ (∀ i w, σ1' !! i = Some w → σ2' !! i = Some w))) v n).
{ eauto. } simpl. clear Hsteps n v1 σ2.
intros v b σ' Hsteps.
eapply ic_adequacy; eauto.
- intros. simpl.
iIntros. iMod (allocate_full_state σ1) as (γh) "[H1 H2]".
iModIntro. iExists γh; rewrite /ownP_full; iFrame.
- iIntros (HICG γh) "Hown".
iApply fupd_ic.
iMod (@special_reg_alloc ⊤ _ ICG_heapG _ l l' ρ τ1) as
(HLogRelG) "[#Hinv #Hll]".
specialize (Hlog HICG HLogRelG).
iPoseProof (Hlog [] [(LitV $ LitLoc l, LitV $ LitLoc l')] with "[]")
as "He".
{ iFrame "#".
rewrite interp_env_cons; iSplit; auto. iApply interp_env_nil. }
iMod (allocate_full_state σ1') as (γh') "[Hfl' Hown']".
iSpecialize ("He" $! γh γh' _ with "[$Hfl']").
simpl in *.
repeat erewrite <- n_closed_subst_head_simpl_singleton;
eauto; try eapply (typed_n_closed [_]); eauto.
simpl.
iModIntro.
iApply ic_wand_r; iSplitR "Hown Hown'"; first iApply "He".
iIntros (v' n'); simpl. iDestruct 1 as (σ2' v2') "(Hrd & Hfl & _)".
iDestruct (heap_included with "[$Hown' $Hfl]") as %Hhi.
iDestruct "Hrd" as %Hrd%rtc_nsteps. destruct Hrd as [n Hrd].
iPureIntro. eexists _, _; split; first eauto using nsteps_rtc, nsteps_det.
destruct Hhi as [Hhi Hvl].
intros i w Hw.
destruct Hhi as [z Hhi].
specialize (Hhi i). rewrite lookup_op in Hhi.
erewrite (lookup_to_heap_Some σ1') in Hhi; eauto.
specialize (Hvl i). erewrite Hhi in Hvl.
revert Hvl Hhi.
case: (z !! i); inversion 1.
unfold to_heap. rewrite lookup_fmap.
case: (σ2' !! i); inversion 1; subst.
apply Excl_inj in H2.
by rewrite (leibniz_equiv _ _ H2).
Qed.
Theorem state_independence e τ ρ τ1 l:
typed [TSTref ρ τ1] e τ →
(∃ σ1 σ2 v, rtc pstep (e.[Lit $ LitLoc l/], σ1) (of_val v, σ2)) →
∀ l' σ1', ∃ σ2' v', rtc pstep (e.[Lit $ LitLoc l'/], σ1') (of_val v', σ2')
∧ (∀ i u, σ1' !! i = Some u → σ2' !! i = Some u).
Proof.
set (Σ := #[ICΣ; PreLogRelΣ]).
intros He (σ1 & σ2 & v & Hrd) l' σ1'.
apply (basic_state_independence Σ e σ1 e σ1' ρ τ1 τ v σ2 l l' He He Hrd).
by intros; eapply binary_fundamental.
Qed.
reduction contextual_refinement IC_adequacy saved_pred ST_Lang_reduction
logrel_shared logrel_ctx_closure.
From iris.algebra Require Import gmap excl dec_agree.
From iris.base_logic Require Import big_op auth.
From iris.proofmode Require Import tactics.
Lemma special_reg_alloc E `{heapG Σ} `{PreLogRelG Σ} l l' ρ
τ :
True ={E}=★ ∃ _ : LogRelG Σ,
reg_inv ★ STRefREL ((@type_to_reg Σ ρ) [])
(⟦τ⟧ []) (LitV (LitLoc l), LitV (LitLoc l')).
Proof.
iIntros "".
iMod (own_alloc (● (∅ : regsUR))) as (γ) "Hregs"; first done.
set (LRG := PreLogRelG_LogRelG γ).
iExists (LRG).
rewrite reg_inv_eq /reg_inv_def /_reg_inv STRefREL_eq /STRefREL_def /=.
iMod (saved_pred_alloc_strong (⟦τ⟧ []) ∅) as (γΦ) "[_ #Hpred]".
iMod (alloc_empty_preds) as (γrel) "HPR".
iMod (Rel_name_alloc _ _ l l' γΦ with "HPR") as "[HPR #Hll]"; first done.
iAssert (Rel_for γrel l l' (⟦τ⟧ [])) as "#HRF";
first by rewrite Rel_for_eq /Rel_for_def /=; iExists _; iFrame "#".
iMod (alloc_empty_bij) as (γbij) "Hbij".
iMod (bij_alloc _ _ l l' with "Hbij") as "[Hbij Hlbij]"; try done.
rewrite union_empty_r_L.
iMod (PREDS_PREDS_include with "HPR") as "[HPR HPRi]".
iMod (own_update with "Hregs") as "Hregs".
{ apply auth_update_alloc.
by apply (alloc_singleton_local_update
_ 1%positive (DecAgree (γbij, γrel))). }
rewrite own_op. iDestruct "Hregs" as "[Hregs How]".
iMod (inv_alloc regIN E
(∃ Rs, REGS Rs ★
[★ map] n ↦ dR ∈ Rs, ∃ R, dR = DecAgree R ★
(∃ g, BIJ (fst R) g ★ PREDS_include (snd R) g))%I
with "[Hregs Hbij HPRi]") as "#Hinv".
- rewrite REGS_eq /REGS_def.
iNext; iExists _; iFrame.
rewrite big_sepM_insert; last done.
iSplitL; last by rewrite big_sepM_empty.
iExists (_, _); iSplit; first done; simpl.
iExists _; iFrame.
- iModIntro. iFrame "#".
iExists l, l', (γbij, γrel); simpl; iFrame; iFrame "#".
iSplit; trivial.
rewrite REG_eq /REG_def type_to_reg_eq /type_to_reg_def /=.
by destruct ρ.
Qed.
Lemma n_closed_subst_head_simpl_singleton e w :
(∀ f, e.[up f] = e) →
e.[of_val w/] = e.[env_subst [w]].
Proof.
intros H1.
rewrite /env_subst. eapply (n_closed_invariant 1); eauto=> /= -[|x] ? //=.
omega.
Qed.
Notation ICG_ST Σ := (@ICG' state ICState_state Σ).
Global Instance ICG_heapG `{H: ICG_ST Σ} : heapG Σ :=
{| heap_invG := IC_invG; heap_inG := ICstate_inG |}.
Lemma basic_state_independence Σ `{ICPreG ST_lang Σ} `{PreLogRelG Σ}
e σ1 e' σ1' ρ τ1 τ v1 σ2 l l' :
[TSTref ρ τ1] ⊢ₜ e : τ → [TSTref ρ τ1] ⊢ₜ e' : τ →
rtc pstep (e.[Lit $ LitLoc l/], σ1) (of_val v1, σ2) →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, [TSTref ρ τ1] ⊨ e ≤log≤ e' : τ) →
(∃ σ2' v2', rtc pstep (e'.[Lit $ LitLoc l'/], σ1') (of_val v2', σ2')
∧ (∀ i w, σ1' !! i = Some w → σ2' !! i = Some w)).
Proof.
intros Hte Hte' Hsteps Hlog.
apply rtc_nsteps in Hsteps. destruct Hsteps as [n Hsteps].
cut (∀ v n σ', nsteps pstep n (e.[Lit $ LitLoc l/], σ1) (of_val v, σ') →
((λ _ _, ∃ σ2' v2',
rtc pstep (e'.[Lit $ LitLoc l'/], σ1') (of_val v2', σ2')
∧ (∀ i w, σ1' !! i = Some w → σ2' !! i = Some w))) v n).
{ eauto. } simpl. clear Hsteps n v1 σ2.
intros v b σ' Hsteps.
eapply ic_adequacy; eauto.
- intros. simpl.
iIntros. iMod (allocate_full_state σ1) as (γh) "[H1 H2]".
iModIntro. iExists γh; rewrite /ownP_full; iFrame.
- iIntros (HICG γh) "Hown".
iApply fupd_ic.
iMod (@special_reg_alloc ⊤ _ ICG_heapG _ l l' ρ τ1) as
(HLogRelG) "[#Hinv #Hll]".
specialize (Hlog HICG HLogRelG).
iPoseProof (Hlog [] [(LitV $ LitLoc l, LitV $ LitLoc l')] with "[]")
as "He".
{ iFrame "#".
rewrite interp_env_cons; iSplit; auto. iApply interp_env_nil. }
iMod (allocate_full_state σ1') as (γh') "[Hfl' Hown']".
iSpecialize ("He" $! γh γh' _ with "[$Hfl']").
simpl in *.
repeat erewrite <- n_closed_subst_head_simpl_singleton;
eauto; try eapply (typed_n_closed [_]); eauto.
simpl.
iModIntro.
iApply ic_wand_r; iSplitR "Hown Hown'"; first iApply "He".
iIntros (v' n'); simpl. iDestruct 1 as (σ2' v2') "(Hrd & Hfl & _)".
iDestruct (heap_included with "[$Hown' $Hfl]") as %Hhi.
iDestruct "Hrd" as %Hrd%rtc_nsteps. destruct Hrd as [n Hrd].
iPureIntro. eexists _, _; split; first eauto using nsteps_rtc, nsteps_det.
destruct Hhi as [Hhi Hvl].
intros i w Hw.
destruct Hhi as [z Hhi].
specialize (Hhi i). rewrite lookup_op in Hhi.
erewrite (lookup_to_heap_Some σ1') in Hhi; eauto.
specialize (Hvl i). erewrite Hhi in Hvl.
revert Hvl Hhi.
case: (z !! i); inversion 1.
unfold to_heap. rewrite lookup_fmap.
case: (σ2' !! i); inversion 1; subst.
apply Excl_inj in H2.
by rewrite (leibniz_equiv _ _ H2).
Qed.
Theorem state_independence e τ ρ τ1 l:
typed [TSTref ρ τ1] e τ →
(∃ σ1 σ2 v, rtc pstep (e.[Lit $ LitLoc l/], σ1) (of_val v, σ2)) →
∀ l' σ1', ∃ σ2' v', rtc pstep (e.[Lit $ LitLoc l'/], σ1') (of_val v', σ2')
∧ (∀ i u, σ1' !! i = Some u → σ2' !! i = Some u).
Proof.
set (Σ := #[ICΣ; PreLogRelΣ]).
intros He (σ1 & σ2 & v & Hrd) l' σ1'.
apply (basic_state_independence Σ e σ1 e σ1' ρ τ1 τ v σ2 l l' He He Hrd).
by intros; eapply binary_fundamental.
Qed.