RunST.logrel
From iris.prelude Require Import tactics.
From iris.proofmode Require Import tactics.
From iris.base_logic Require Export invariants.
From RunST Require Import lang IC rules typing ST_Lang_reduction
saved_pred regions reduction logrel_shared.
From iris.algebra Require Import list.
Import uPred.
Set Bullet Behavior "Strict Subproofs".
(* HACK: move somewhere else *)
Ltac auto_equiv ::=
(* Deal with "pointwise_relation" *)
repeat lazymatch goal with
| |- pointwise_relation _ _ _ _ => intros ?
end;
(* Normalize away equalities. *)
repeat match goal with
| H : _ ≡{_}≡ _ |- _ => apply (timeless_iff _ _) in H
| _ => progress simplify_eq
end;
(* repeatedly apply congruence lemmas and use the equalities in the hypotheses. *)
try (f_equiv; fast_done || auto_equiv).
From iris.proofmode Require Import tactics.
From iris.base_logic Require Export invariants.
From RunST Require Import lang IC rules typing ST_Lang_reduction
saved_pred regions reduction logrel_shared.
From iris.algebra Require Import list.
Import uPred.
Set Bullet Behavior "Strict Subproofs".
(* HACK: move somewhere else *)
Ltac auto_equiv ::=
(* Deal with "pointwise_relation" *)
repeat lazymatch goal with
| |- pointwise_relation _ _ _ _ => intros ?
end;
(* Normalize away equalities. *)
repeat match goal with
| H : _ ≡{_}≡ _ |- _ => apply (timeless_iff _ _) in H
| _ => progress simplify_eq
end;
(* repeatedly apply congruence lemmas and use the equalities in the hypotheses. *)
try (f_equiv; fast_done || auto_equiv).
interp : is a unary logical relation.
Section logrel.
Context `{heapG Σ, LogRelG Σ}.
Notation Cx := (prodC ((prodC valC valC) -n> iProp Σ) (leibnizC positive)).
Notation D := (prodC valC valC -n> iProp Σ).
Implicit Types τi : D.
Implicit Types Δ : listC Cx.
Implicit Types interp : listC Cx → D.
Notation prim_step' := (pstep' det_lang).
Notation steps' := (rtc prim_step').
Definition interp_expr (τi : listC Cx -n> D) (Δ : listC Cx)
(ee : expr * expr) : iProp Σ := (∀ γh γh' σ1',
ownP_full γh' σ1' -★
IC ee.1 {{ v, ∃ σ2' v',
(■ rtc (prim_step') (ee.2, σ1') (of_val v', σ2'))
★ ownP_full γh' σ2' ★ τi Δ (v, v') }}[γh])%I.
Global Instance interp_expr_ne n :
Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr.
Proof.
intros ? ? Heq ? ? Heq' ? ? Heq''. rewrite /interp_expr /=; subst.
repeat apply forall_ne =>?; apply wand_ne; auto.
apply ic_ne => ? ? ? Heq''.
repeat apply exist_ne =>? /=; apply sep_ne; auto.
by rewrite Heq Heq'.
Qed.
Program Definition interp_TST ρ (τi : listC Cx -n> D) : listC Cx -n> D :=
λne Δ vv,
(□ ∀ γh γh' σ1',
ownP_full γh' σ1' ★ Region (type_to_reg ρ Δ) γh γh' -★
IC RunST (of_val (vv.1)) {{ v, ∃ σ2' v',
(■ (steps' (RunST (of_val (vv.2)), σ1') (of_val v', σ2')))
★ ownP_full γh' σ2' ★ τi Δ (v, v')
★ Region (type_to_reg ρ Δ) γh γh' }}[γh])%I.
Solve Obligations with solve_proper.
Next Obligation.
Proof.
intros ? ? ? ? ? Heq ?; simpl. apply always_ne.
repeat apply forall_ne =>?. apply wand_ne.
- apply sep_ne; auto. by rewrite Heq.
- apply ic_ne => ? ? ? Heq'. repeat apply exist_ne =>?.
by rewrite Heq.
Qed.
Fixpoint interp (τ : type) : listC Cx -n> D :=
match τ return _ with
| TUnit => interp_unit
| TInt => interp_int
| TBool => interp_bool
| TProd τ1 τ2 => interp_prod (interp τ1) (interp τ2)
| TSum τ1 τ2 => interp_sum (interp τ1) (interp τ2)
| TArrow τ1 τ2 =>
interp_arrow interp_expr interp_expr_ne (interp τ1) (interp τ2)
| TVar x => ctx_lookup x
| TForall τ' => interp_forall interp_expr interp_expr_ne (interp τ')
| TRec τ' => interp_rec τ' (interp τ')
| TSTref ρ τ' => interp_ref ρ (interp τ')
| TST ρ τ' => interp_TST ρ (interp τ')
end.
Notation "⟦ τ ⟧" := (interp τ).
Notation "⟦ Γ ⟧*" := (interp_env interp Γ).
Global Instance interp_persistent τ Δ vv :
env_PersistentP Δ → PersistentP (⟦ τ ⟧ Δ vv).
Proof.
revert vv Δ; induction τ=> vv Δ HΔ; simpl; try apply _.
rewrite /PersistentP /interp_rec fixpoint_unfold /interp_rec1 /=.
by apply always_intro'.
Qed.
Lemma type_to_reg_subst_up Δ1 Δ2 τ τ' :
(type_to_reg τ)
(Δ1 ++ (⟦ τ' ⟧ Δ2,
(type_to_reg (τ'.[ren (+length Δ1)]) (Δ1 ++ Δ2))) :: Δ2) =
(type_to_reg τ.[upn (length Δ1) (τ' .: ids)]) (Δ1 ++ Δ2).
Proof.
rewrite type_to_reg_eq /type_to_reg_def /=.
destruct τ; trivial.
asimpl. rewrite iter_up.
destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r;
last by revert Hl; move: (length Δ1) => ? ?; lia.
destruct (x - length Δ1) as [|n] eqn:?; simpl; rewrite Heqn.
{ by asimpl. }
simpl. rewrite !lookup_app_r; first do 2 f_equiv; trivial.
revert Hl Heqn; move: (length Δ1) => ? ? ?; lia.
move: (length Δ1) => ?; lia.
Qed.
Lemma interp_weaken Δ1 Π Δ2 τ :
⟦ τ.[upn (length Δ1) (ren (+ length Π))] ⟧ (Δ1 ++ Π ++ Δ2)
≡ ⟦ τ ⟧ (Δ1 ++ Δ2).
Proof.
revert Δ1 Π Δ2. induction τ=> Δ1 Π Δ2; simpl; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- apply fixpoint_proper=> τi ww /=.
properness; auto. asimpl. rewrite -(IHτ (_ :: _)).
rewrite /= (type_to_reg_weaken Δ1 Π Δ2 (TRec τ)); auto.
- rewrite iter_up. destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r; first do 2 f_equiv; trivial.
all: revert Hl; move: (length Δ1) (length Π) => ? ? ?; lia.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply (IHτ (_ :: _)).
- intros ww; simpl. rewrite IHτ2. by rewrite type_to_reg_weaken.
- intros ww; simpl; properness; auto; try apply IHτ2;
by rewrite type_to_reg_weaken.
Qed.
Lemma interp_subst_up Δ1 Δ2 τ τ' :
⟦ τ ⟧
(Δ1 ++ (interp τ' Δ2,
(type_to_reg (τ'.[ren (+length Δ1)]) (Δ1 ++ Δ2))) :: Δ2)
≡ ⟦ τ.[upn (length Δ1) (τ' .: ids)] ⟧ (Δ1 ++ Δ2).
Proof.
revert Δ1 Δ2; induction τ=> Δ1 Δ2; simpl; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- apply fixpoint_proper=> τi ww /=.
properness; auto. rewrite -(IHτ (_ :: _)). asimpl.
rewrite (type_to_reg_subst_up Δ1 Δ2 (TRec τ) τ'). asimpl.
pose proof (type_to_reg_weaken
[] [(τi, (type_to_reg
(TRec τ.[upn (S (length Δ1)) (τ' .: ids)])) (Δ1 ++ Δ2))]
(Δ1 ++ Δ2) (τ'.[ren (+length Δ1)])) as Hlem.
asimpl in Hlem. by rewrite Hlem.
- rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r;
last by revert Hl; move: (length Δ1) => ? ?; lia.
destruct (x - length Δ1) as [|n] eqn:?; simpl; rewrite Heqn.
{ symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). }
simpl. rewrite !lookup_app_r; first do 2 f_equiv; trivial.
revert Hl Heqn; move: (length Δ1) => ? ? ?; lia.
move: (length Δ1) => ?; lia.
- unfold interp_expr.
intros ww; simpl.
apply always_proper.
apply forall_proper => τi.
apply forall_proper => n.
properness; auto.
rewrite -(IHτ (_ :: _)). asimpl.
pose proof (type_to_reg_weaken [] [(τi, n)]
(Δ1 ++ Δ2) (τ'.[ren (+length Δ1)])) as Hlem.
asimpl in Hlem. by rewrite Hlem.
- intros ww; simpl; properness; auto. rewrite IHτ2.
by rewrite type_to_reg_subst_up.
- intros ww; simpl; properness; auto; try apply IHτ2;
by rewrite type_to_reg_subst_up.
Qed.
Lemma interp_subst Δ2 τ τ' :
⟦ τ ⟧ ((⟦ τ' ⟧ Δ2, type_to_reg τ' Δ2) :: Δ2) ≡ ⟦ τ.[τ'/] ⟧ Δ2.
Proof. by rewrite -(interp_subst_up []); asimpl. Qed.
Lemma interp_env_ren p Δ (Γ : list type) vvs τi :
⟦ subst (ren (+1)) <$> Γ ⟧* ((τi, p) :: Δ) vvs ⊣⊢ ⟦ Γ ⟧* Δ vvs.
Proof.
apply sep_proper; [apply pure_proper; by rewrite fmap_length|].
revert Δ vvs τi; induction Γ=> Δ [|v vs] τi; csimpl; auto.
apply sep_proper; auto. apply (interp_weaken [] [(τi, p)] Δ).
Qed.
End logrel.
Typeclasses Opaque interp_env.
Notation "⟦ τ ⟧" := (interp τ).
Notation "⟦ τ ⟧ₑ" := (interp_expr (interp τ)).
Notation "⟦ Γ ⟧*" := (interp_env interp Γ).
Context `{heapG Σ, LogRelG Σ}.
Notation Cx := (prodC ((prodC valC valC) -n> iProp Σ) (leibnizC positive)).
Notation D := (prodC valC valC -n> iProp Σ).
Implicit Types τi : D.
Implicit Types Δ : listC Cx.
Implicit Types interp : listC Cx → D.
Notation prim_step' := (pstep' det_lang).
Notation steps' := (rtc prim_step').
Definition interp_expr (τi : listC Cx -n> D) (Δ : listC Cx)
(ee : expr * expr) : iProp Σ := (∀ γh γh' σ1',
ownP_full γh' σ1' -★
IC ee.1 {{ v, ∃ σ2' v',
(■ rtc (prim_step') (ee.2, σ1') (of_val v', σ2'))
★ ownP_full γh' σ2' ★ τi Δ (v, v') }}[γh])%I.
Global Instance interp_expr_ne n :
Proper (dist n ==> dist n ==> (=) ==> dist n) interp_expr.
Proof.
intros ? ? Heq ? ? Heq' ? ? Heq''. rewrite /interp_expr /=; subst.
repeat apply forall_ne =>?; apply wand_ne; auto.
apply ic_ne => ? ? ? Heq''.
repeat apply exist_ne =>? /=; apply sep_ne; auto.
by rewrite Heq Heq'.
Qed.
Program Definition interp_TST ρ (τi : listC Cx -n> D) : listC Cx -n> D :=
λne Δ vv,
(□ ∀ γh γh' σ1',
ownP_full γh' σ1' ★ Region (type_to_reg ρ Δ) γh γh' -★
IC RunST (of_val (vv.1)) {{ v, ∃ σ2' v',
(■ (steps' (RunST (of_val (vv.2)), σ1') (of_val v', σ2')))
★ ownP_full γh' σ2' ★ τi Δ (v, v')
★ Region (type_to_reg ρ Δ) γh γh' }}[γh])%I.
Solve Obligations with solve_proper.
Next Obligation.
Proof.
intros ? ? ? ? ? Heq ?; simpl. apply always_ne.
repeat apply forall_ne =>?. apply wand_ne.
- apply sep_ne; auto. by rewrite Heq.
- apply ic_ne => ? ? ? Heq'. repeat apply exist_ne =>?.
by rewrite Heq.
Qed.
Fixpoint interp (τ : type) : listC Cx -n> D :=
match τ return _ with
| TUnit => interp_unit
| TInt => interp_int
| TBool => interp_bool
| TProd τ1 τ2 => interp_prod (interp τ1) (interp τ2)
| TSum τ1 τ2 => interp_sum (interp τ1) (interp τ2)
| TArrow τ1 τ2 =>
interp_arrow interp_expr interp_expr_ne (interp τ1) (interp τ2)
| TVar x => ctx_lookup x
| TForall τ' => interp_forall interp_expr interp_expr_ne (interp τ')
| TRec τ' => interp_rec τ' (interp τ')
| TSTref ρ τ' => interp_ref ρ (interp τ')
| TST ρ τ' => interp_TST ρ (interp τ')
end.
Notation "⟦ τ ⟧" := (interp τ).
Notation "⟦ Γ ⟧*" := (interp_env interp Γ).
Global Instance interp_persistent τ Δ vv :
env_PersistentP Δ → PersistentP (⟦ τ ⟧ Δ vv).
Proof.
revert vv Δ; induction τ=> vv Δ HΔ; simpl; try apply _.
rewrite /PersistentP /interp_rec fixpoint_unfold /interp_rec1 /=.
by apply always_intro'.
Qed.
Lemma type_to_reg_subst_up Δ1 Δ2 τ τ' :
(type_to_reg τ)
(Δ1 ++ (⟦ τ' ⟧ Δ2,
(type_to_reg (τ'.[ren (+length Δ1)]) (Δ1 ++ Δ2))) :: Δ2) =
(type_to_reg τ.[upn (length Δ1) (τ' .: ids)]) (Δ1 ++ Δ2).
Proof.
rewrite type_to_reg_eq /type_to_reg_def /=.
destruct τ; trivial.
asimpl. rewrite iter_up.
destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r;
last by revert Hl; move: (length Δ1) => ? ?; lia.
destruct (x - length Δ1) as [|n] eqn:?; simpl; rewrite Heqn.
{ by asimpl. }
simpl. rewrite !lookup_app_r; first do 2 f_equiv; trivial.
revert Hl Heqn; move: (length Δ1) => ? ? ?; lia.
move: (length Δ1) => ?; lia.
Qed.
Lemma interp_weaken Δ1 Π Δ2 τ :
⟦ τ.[upn (length Δ1) (ren (+ length Π))] ⟧ (Δ1 ++ Π ++ Δ2)
≡ ⟦ τ ⟧ (Δ1 ++ Δ2).
Proof.
revert Δ1 Π Δ2. induction τ=> Δ1 Π Δ2; simpl; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- apply fixpoint_proper=> τi ww /=.
properness; auto. asimpl. rewrite -(IHτ (_ :: _)).
rewrite /= (type_to_reg_weaken Δ1 Π Δ2 (TRec τ)); auto.
- rewrite iter_up. destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r; first do 2 f_equiv; trivial.
all: revert Hl; move: (length Δ1) (length Π) => ? ? ?; lia.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply (IHτ (_ :: _)).
- intros ww; simpl. rewrite IHτ2. by rewrite type_to_reg_weaken.
- intros ww; simpl; properness; auto; try apply IHτ2;
by rewrite type_to_reg_weaken.
Qed.
Lemma interp_subst_up Δ1 Δ2 τ τ' :
⟦ τ ⟧
(Δ1 ++ (interp τ' Δ2,
(type_to_reg (τ'.[ren (+length Δ1)]) (Δ1 ++ Δ2))) :: Δ2)
≡ ⟦ τ.[upn (length Δ1) (τ' .: ids)] ⟧ (Δ1 ++ Δ2).
Proof.
revert Δ1 Δ2; induction τ=> Δ1 Δ2; simpl; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- unfold interp_expr.
intros ww; simpl; properness; auto. by apply IHτ1. by apply IHτ2.
- apply fixpoint_proper=> τi ww /=.
properness; auto. rewrite -(IHτ (_ :: _)). asimpl.
rewrite (type_to_reg_subst_up Δ1 Δ2 (TRec τ) τ'). asimpl.
pose proof (type_to_reg_weaken
[] [(τi, (type_to_reg
(TRec τ.[upn (S (length Δ1)) (τ' .: ids)])) (Δ1 ++ Δ2))]
(Δ1 ++ Δ2) (τ'.[ren (+length Δ1)])) as Hlem.
asimpl in Hlem. by rewrite Hlem.
- rewrite iter_up; destruct lt_dec as [Hl | Hl]; simpl.
{ by rewrite !lookup_app_l. }
rewrite !lookup_app_r;
last by revert Hl; move: (length Δ1) => ? ?; lia.
destruct (x - length Δ1) as [|n] eqn:?; simpl; rewrite Heqn.
{ symmetry. asimpl. apply (interp_weaken [] Δ1 Δ2 τ'). }
simpl. rewrite !lookup_app_r; first do 2 f_equiv; trivial.
revert Hl Heqn; move: (length Δ1) => ? ? ?; lia.
move: (length Δ1) => ?; lia.
- unfold interp_expr.
intros ww; simpl.
apply always_proper.
apply forall_proper => τi.
apply forall_proper => n.
properness; auto.
rewrite -(IHτ (_ :: _)). asimpl.
pose proof (type_to_reg_weaken [] [(τi, n)]
(Δ1 ++ Δ2) (τ'.[ren (+length Δ1)])) as Hlem.
asimpl in Hlem. by rewrite Hlem.
- intros ww; simpl; properness; auto. rewrite IHτ2.
by rewrite type_to_reg_subst_up.
- intros ww; simpl; properness; auto; try apply IHτ2;
by rewrite type_to_reg_subst_up.
Qed.
Lemma interp_subst Δ2 τ τ' :
⟦ τ ⟧ ((⟦ τ' ⟧ Δ2, type_to_reg τ' Δ2) :: Δ2) ≡ ⟦ τ.[τ'/] ⟧ Δ2.
Proof. by rewrite -(interp_subst_up []); asimpl. Qed.
Lemma interp_env_ren p Δ (Γ : list type) vvs τi :
⟦ subst (ren (+1)) <$> Γ ⟧* ((τi, p) :: Δ) vvs ⊣⊢ ⟦ Γ ⟧* Δ vvs.
Proof.
apply sep_proper; [apply pure_proper; by rewrite fmap_length|].
revert Δ vvs τi; induction Γ=> Δ [|v vs] τi; csimpl; auto.
apply sep_proper; auto. apply (interp_weaken [] [(τi, p)] Δ).
Qed.
End logrel.
Typeclasses Opaque interp_env.
Notation "⟦ τ ⟧" := (interp τ).
Notation "⟦ τ ⟧ₑ" := (interp_expr (interp τ)).
Notation "⟦ Γ ⟧*" := (interp_env interp Γ).