diff --git a/theories/logic/model.v b/theories/logic/model.v index 7dcfb674588888902085a7780d002f63e79ca3bc..5dfc08bf43fd47c1556740e3ab0b4d40e095997a 100644 --- a/theories/logic/model.v +++ b/theories/logic/model.v @@ -63,6 +63,10 @@ Section semtypes. Proper ((=) ==> (=) ==> dist n ==> dist n) (interp_expr E). Proof. solve_proper. Qed. + Global Instance interp_expr_proper E e e' : + Proper ((≡) ==> (≡)) (interp_expr E e e'). + Proof. apply ne_proper=>n. by apply interp_expr_ne. Qed. + Definition lty2_unit : lty2 := Lty2 (λ w1 w2, ⌜ w1 = #() ∧ w2 = #() âŒ%I). Definition lty2_bool : lty2 := Lty2 (λ w1 w2, ∃ b : bool, ⌜ w1 = #b ∧ w2 = #b âŒ)%I. Definition lty2_int : lty2 := Lty2 (λ w1 w2, ∃ n : Z, ⌜ w1 = #n ∧ w2 = #n âŒ)%I. @@ -119,6 +123,9 @@ Section semtypes. apply lty2_car_ne; eauto. Qed. + Lemma lty_rec_unfold (C : lty2C -n> lty2C) : lty2_rec C ≡ lty2_rec1 C (lty2_rec C). + Proof. apply fixpoint_unfold. Qed. + End semtypes. (* Nice notations *) diff --git a/theories/typing/fundamental.v b/theories/typing/fundamental.v index 5e1137cf440d96749a9cfddbe5be4290b8cbf749..fe5b2c6649657ee73cad44813a0a9ed0d068cc03 100644 --- a/theories/typing/fundamental.v +++ b/theories/typing/fundamental.v @@ -6,6 +6,7 @@ From reloc.logic.proofmode Require Export tactics spec_tactics. From reloc.typing Require Export interp. From iris.proofmode Require Export tactics. +From Autosubst Require Import Autosubst. Section fundamental. Context `{relocG Σ}. @@ -381,9 +382,6 @@ Section fundamental. destruct op; inversion Hopv'; simplify_eq/=; eauto. Qed. - (* TODO *) - - From Autosubst Require Import Autosubst. Lemma bin_log_related_tapp' Δ Γ e e' Ï„ Ï„' : ({Δ;Γ} ⊨ e ≤log≤ e' : TForall Ï„) -∗ {Δ;Γ} ⊨ (TApp e) ≤log≤ (TApp e') : Ï„.[Ï„'/]. @@ -394,52 +392,53 @@ Section fundamental. rel_bind_ap e e' "IH" v v' "IH". iDestruct ("IH" $! (interp Ï„' Δ)) as "#IH". iApply refines_ret'; auto. - admit. - Admitted. + by rewrite -interp_subst. + Qed. -(* - Lemma bin_log_related_tapp (Ï„i : D) Δ Γ e e' Ï„ : - {Δ;Γ} ⊨ e ≤log≤ e' : TForall Ï„ -∗ - {Ï„i::Δ;⤉Γ} ⊨ TApp e ≤log≤ TApp e' : Ï„. + Lemma bin_log_related_tapp (Ï„i : lty2) Δ Γ e e' Ï„ : + ({Δ;Γ} ⊨ e ≤log≤ e' : TForall Ï„) -∗ + {Ï„i::Δ;⤉Γ} ⊨ (TApp e) ≤log≤ (TApp e') : Ï„. Proof. - rewrite bin_log_related_eq. + rewrite /bin_log_related refines_eq. iIntros "IH". iIntros (vvs Ï) "#Hs #HΓ"; iIntros (j K) "Hj /=". - wp_bind (env_subst _ e). - tp_bind j (env_subst _ e'). + iModIntro. wp_bind (subst_map _ e). + tp_bind j (subst_map _ e'). iSpecialize ("IH" with "Hs [HΓ] Hj"). - { by rewrite interp_env_ren. } - iMod "IH" as "IH /=". iModIntro. + { by rewrite interp_ren. } + iMod "IH" as "IH /=". iApply (wp_wand with "IH"). iIntros (v). iDestruct 1 as (v') "[Hj #IH]". iMod ("IH" $! Ï„i with "Hj"); auto. Qed. - Lemma bin_log_related_fold Δ Γ e e' Ï„ : - {Δ;Γ} ⊨ e ≤log≤ e' : Ï„.[(TRec Ï„)/] -∗ - {Δ;Γ} ⊨ Fold e ≤log≤ Fold e' : TRec Ï„. + Lemma bin_log_related_unfold Δ Γ e e' Ï„ : + ({Δ;Γ} ⊨ e ≤log≤ e' : TRec Ï„) -∗ + {Δ;Γ} ⊨ rec_unfold e ≤log≤ rec_unfold e' : Ï„.[(TRec Ï„)/]. Proof. iIntros "IH". rel_bind_ap e e' "IH" v v' "IH". - value_case. - rewrite fixpoint_unfold /= -interp_subst. - iExists (_, _); eauto. + iEval (rewrite lty_rec_unfold /lty2_car /=) in "IH". + change (lty2_rec _) with (interp (TRec Ï„) Δ). + unfold rec_unfold. unlock. (* TODO WHY?????? *) + rel_pure_l. rel_pure_r. + value_case. by rewrite -interp_subst. Qed. - Lemma bin_log_related_unfold Δ Γ e e' Ï„ : - {Δ;Γ} ⊨ e ≤log≤ e' : TRec Ï„ -∗ - {Δ;Γ} ⊨ Unfold e ≤log≤ Unfold e' : Ï„.[(TRec Ï„)/]. + Lemma bin_log_related_fold Δ Γ e e' Ï„ : + ({Δ;Γ} ⊨ e ≤log≤ e' : Ï„.[(TRec Ï„)/]) -∗ + {Δ;Γ} ⊨ e ≤log≤ e' : TRec Ï„. Proof. iIntros "IH". rel_bind_ap e e' "IH" v v' "IH". - rewrite /= fixpoint_unfold /=. - change (fixpoint _) with (interp ⊤ (TRec Ï„) Δ). - iDestruct "IH" as ([w w']) "#[% Hiz]"; simplify_eq/=. - rel_unfold_l. - rel_unfold_r. - value_case. by rewrite -interp_subst. + value_case. + iModIntro. + iEval (rewrite lty_rec_unfold /lty2_car /=). + change (lty2_rec _) with (interp (TRec Ï„) Δ). + by rewrite -interp_subst. Qed. +(* Lemma bin_log_related_pack' Δ Γ e e' Ï„ Ï„' : {Δ;Γ} ⊨ e ≤log≤ e' : Ï„.[Ï„'/] -∗ {Δ;Γ} ⊨ Pack e ≤log≤ Pack e' : TExists Ï„. @@ -501,7 +500,7 @@ Section fundamental. rewrite (interp_weaken [] [Ï„i] Δ _ Ï„2) /=. eauto. Qed. - +*) Theorem binary_fundamental Δ Γ e Ï„ : Γ ⊢ₜ e : Ï„ → ({Δ;Γ} ⊨ e ≤log≤ e : Ï„)%I. @@ -521,30 +520,19 @@ Section fundamental. - by iApply bin_log_related_injr. - by iApply (bin_log_related_case with "[] []"). - by iApply (bin_log_related_if with "[] []"). - - assert (Closed (x :b: f :b: dom (gset string) Γ) e). - { apply typed_X_closed in Ht. - rewrite !dom_insert_binder in Ht. - revert Ht. destruct x,f; cbn-[union]; - (* TODO: why is simple rewriting is not sufficient here? *) - erewrite ?(left_id ∅); eauto. - all: apply _. } - iApply (bin_log_related_rec with "[]"); eauto. + - iApply (bin_log_related_rec with "[]"); eauto. - by iApply (bin_log_related_app with "[] []"). - - assert (Closed (dom _ Γ) e). - { apply typed_X_closed in Ht. - pose (K:=(dom_fmap (asubst (ren (+1))) Γ (D:=stringset))). - fold_leibniz. by rewrite -K. } - iApply bin_log_related_tlam; eauto. + - iApply bin_log_related_tlam; eauto. - by iApply bin_log_related_tapp'. - by iApply bin_log_related_fold. - by iApply bin_log_related_unfold. - - by iApply bin_log_related_pack'. - - iApply (bin_log_related_unpack with "[]"); eauto. + (* - by iApply bin_log_related_pack'. *) + (* - iApply (bin_log_related_unpack with "[]"); eauto. *) - by iApply bin_log_related_fork. - by iApply bin_log_related_alloc. - by iApply bin_log_related_load. - by iApply (bin_log_related_store with "[]"). - by iApply (bin_log_related_CAS with "[] []"). Qed. -*) + End fundamental. diff --git a/theories/typing/interp.v b/theories/typing/interp.v index 2d79b913f59343844cf83e0caea7af3fdd8f4833..00e4fa56cfb8b15d9a8b40376718740bad6a1c07 100644 --- a/theories/typing/interp.v +++ b/theories/typing/interp.v @@ -169,4 +169,9 @@ Section interp_ren. destruct (Γ !! x); auto; simpl. f_equiv. symmetry. apply (interp_ren_up []). Qed. + + Lemma interp_subst Δ2 Ï„ Ï„' : + interp Ï„ (interp Ï„' Δ2 :: Δ2) ≡ interp (Ï„.[Ï„'/]) Δ2. + Proof. Admitted. + End interp_ren. diff --git a/theories/typing/types.v b/theories/typing/types.v index 2dad64f3bd4a107fc390198d8b79a423dc5d6fd6..450dc50715530cfde242b5d874c3948489168fc7 100644 --- a/theories/typing/types.v +++ b/theories/typing/types.v @@ -116,7 +116,7 @@ Inductive typed (Γ : stringmap type) : expr → type → Prop := | If_typed e0 e1 e2 Ï„ : Γ ⊢ₜ e0 : TBool → Γ ⊢ₜ e1 : Ï„ → Γ ⊢ₜ e2 : Ï„ → Γ ⊢ₜ If e0 e1 e2 : Ï„ | Rec_typed f x e Ï„1 Ï„2 : - <[x:=Ï„1]>(<[f:=TArrow Ï„1 Ï„2]>Γ) ⊢ₜ e : Ï„2 → + <[f:=TArrow Ï„1 Ï„2]>(<[x:=Ï„1]>Γ) ⊢ₜ e : Ï„2 → Γ ⊢ₜ Rec f x e : TArrow Ï„1 Ï„2 | App_typed e1 e2 Ï„1 Ï„2 : Γ ⊢ₜ e1 : TArrow Ï„1 Ï„2 → Γ ⊢ₜ e2 : Ï„1 → Γ ⊢ₜ App e1 e2 : Ï„2 @@ -127,16 +127,17 @@ Inductive typed (Γ : stringmap type) : expr → type → Prop := Γ ⊢ₜ e #() : Ï„.[Ï„'/] | TFold e Ï„ : Γ ⊢ₜ e : Ï„.[TRec Ï„/] → Γ ⊢ₜe : TRec Ï„ | TUnfold e Ï„ : Γ ⊢ₜ e : TRec Ï„ → Γ ⊢ₜ rec_unfold e : Ï„.[TRec Ï„/] - | TPack e Ï„ Ï„' : Γ ⊢ₜ e : Ï„.[Ï„'/] → Γ ⊢ₜ e : TExists Ï„ - | TUnpack e1 e2 Ï„ Ï„2 : Γ ⊢ₜ e1 : TExists Ï„ → - (subst (ren (+1%nat)) <$> Γ) ⊢ₜ e2 : TArrow Ï„ (subst (ren (+1%nat)) Ï„2) → - Γ ⊢ₜ unpack e1 e2 : Ï„2 + (* | TPack e Ï„ Ï„' : Γ ⊢ₜ e : Ï„.[Ï„'/] → Γ ⊢ₜ e : TExists Ï„ *) + (* | TUnpack e1 e2 Ï„ Ï„2 : Γ ⊢ₜ e1 : TExists Ï„ → *) + (* (subst (ren (+1%nat)) <$> Γ) ⊢ₜ e2 : TArrow Ï„ (subst (ren (+1%nat)) Ï„2) → *) + (* Γ ⊢ₜ unpack e1 e2 : Ï„2 *) | TFork e : Γ ⊢ₜ e : TUnit → Γ ⊢ₜ Fork e : TUnit | TAlloc e Ï„ : Γ ⊢ₜ e : Ï„ → Γ ⊢ₜ Alloc e : Tref Ï„ | TLoad e Ï„ : Γ ⊢ₜ e : Tref Ï„ → Γ ⊢ₜ Load e : Ï„ | TStore e e' Ï„ : Γ ⊢ₜ e : Tref Ï„ → Γ ⊢ₜ e' : Ï„ → Γ ⊢ₜ Store e e' : TUnit | TCAS e1 e2 e3 Ï„ : - EqType Ï„ → Γ ⊢ₜ e1 : Tref Ï„ → Γ ⊢ₜ e2 : Ï„ → Γ ⊢ₜ e3 : Ï„ → + EqType Ï„ → UnboxedType Ï„ → + Γ ⊢ₜ e1 : Tref Ï„ → Γ ⊢ₜ e2 : Ï„ → Γ ⊢ₜ e3 : Ï„ → Γ ⊢ₜ CAS e1 e2 e3 : TBool where "Γ ⊢ₜ e : Ï„" := (typed Γ e Ï„).