diff --git a/theories/logic/compatibility.v b/theories/logic/compatibility.v index 2cb61fb44de178a0582e657f130208acd7837d7a..3011969923459fe23d81d1d862df1a822c49cae3 100644 --- a/theories/logic/compatibility.v +++ b/theories/logic/compatibility.v @@ -124,5 +124,70 @@ Section compatibility. value_case. Qed. + Lemma refines_cmpxchg_ref A e1 e2 e3 e1' e2' e3' : + (REL e1 << e1' : lrel_ref (lrel_ref A)) -∗ + (REL e2 << e2' : lrel_ref A) -∗ + (REL e3 << e3' : lrel_ref A) -∗ + REL (CmpXchg e1 e2 e3) << (CmpXchg e1' e2' e3') : lrel_prod (lrel_ref A) lrel_bool. + Proof. + iIntros "IH1 IH2 IH3". + rel_bind_l e3; rel_bind_r e3'. + iApply (refines_bind with "IH3"). + iIntros (v3 v3') "#IH3". + rel_bind_l e2; rel_bind_r e2'. + iApply (refines_bind with "IH2"). + iIntros (v2 v2') "#IH2". + rel_bind_l e1; rel_bind_r e1'. + iApply (refines_bind with "IH1"). + iIntros (v1 v1') "#IH1 /=". + iPoseProof "IH1" as "IH1'". + iDestruct "IH1" as (l1 l2) "(% & % & Hinv)"; simplify_eq/=. + rewrite {2}/lrel_car /=. + iDestruct "IH2" as (r1 r2 -> ->) "Hr". + (* iDestruct "IH3" as (n1 n2 -> ->) "Hn". *) + rel_cmpxchg_l_atomic. + iInv (relocN .@ "ref" .@ (l1,l2)) as (v1 v2) "[Hl1 [>Hl2 #Hv]]" "Hclose". + iModIntro. iExists _; iFrame. simpl. + destruct (decide ((v1, v2) = (#r1, #r2))); simplify_eq/=. + - iSplitR; first by (iIntros (?); contradiction). + iIntros (?). iNext. iIntros "Hv1". + rel_cmpxchg_suc_r. + iMod ("Hclose" with "[-]"). + { iNext. iExists _, _. by iFrame. } + rel_values. iExists _, _, _, _. do 2 (iSplitL; first done). + iFrame "Hv". iExists _. done. + - iSplit; iIntros (?); simplify_eq/=; iNext; iIntros "Hl1". + + destruct (decide (v2 = #r2)); simplify_eq/=. + * rewrite {5}/lrel_car. simpl. + iDestruct "Hv" as (r1' r2' ? ?) "#Hv". simplify_eq/=. + destruct (decide ((l1, l2) = (r1, r2'))); simplify_eq/=. + { iInv (relocN.@"ref".@(r1', r2')) as (? ?) "(Hr1 & >Hr2 & Hrr)" "Hcl". + iExFalso. by iDestruct (mapstoS_valid_2 with "Hr2 Hl2") as %[]. } + destruct (decide ((l1, l2) = (r1', r2'))); simplify_eq/=. + ++ assert (r1' ≠r1) by naive_solver. + iInv (relocN.@"ref".@(r1, r2')) as (? ?) "(Hr1 & >Hr2 & Hrr)" "Hcl". + iExFalso. by iDestruct (mapstoS_valid_2 with "Hr2 Hl2") as %[]. + ++ iInv (relocN.@"ref".@(r1, r2')) as (? ?) "(Hr1 & >Hr2 & Hrr)" "Hcl". + iInv (relocN.@"ref".@(r1', r2')) as (? ?) "(Hr1' & >Hr2' & Hrr')" "Hcl'". + iExFalso. by iDestruct (mapstoS_valid_2 with "Hr2 Hr2'") as %[]. + * rel_cmpxchg_fail_r. + iMod ("Hclose" with "[-]"). + { iNext; iExists _, _; by iFrame. } + rel_values. iModIntro. iExists _,_,_,_. + repeat iSplit; eauto. + + assert (v2 ≠#r2) by naive_solver. + rewrite {5}/lrel_car. simpl. + iDestruct "Hv" as (r1' r2' ? ?) "#Hv". simplify_eq/=. + destruct (decide ((l1, l2) = (r1', r2'))); simplify_eq/=. + { iInv (relocN.@"ref".@(r1', r2)) as (? ?) "(>Hr1 & >Hr2 & Hrr)" "Hcl". + iExFalso. by iDestruct (mapsto_valid_2 with "Hr1 Hl1") as %[]. } + destruct (decide ((l1, l2) = (r1', r2))); simplify_eq/=. + { iInv (relocN.@"ref".@(r1', r2')) as (? ?) "(>Hr1 & >Hr2 & Hrr)" "Hcl". + iExFalso. by iDestruct (mapsto_valid_2 with "Hr1 Hl1") as %[]. } + iInv (relocN.@"ref".@(r1', r2)) as (? ?) "(>Hr1 & >Hr2 & Hrr)" "Hcl". + iInv (relocN.@"ref".@(r1', r2')) as (? ?) "(>Hr1' & >Hr2' & Hrr')" "Hcl'". + iExFalso. by iDestruct (mapsto_valid_2 with "Hr1 Hr1'") as %[]. + Qed. + End compatibility. diff --git a/theories/typing/contextual_refinement.v b/theories/typing/contextual_refinement.v index bbed1775e30441f8eb77558d09ffebb7b28a186b..89969de51069904e9222b16817d89139ee0b3a66 100644 --- a/theories/typing/contextual_refinement.v +++ b/theories/typing/contextual_refinement.v @@ -138,12 +138,14 @@ Inductive typed_ctx_item : typed Γ e1 TBool → binop_bool_res_type op = Some Ï„ → typed_ctx_item (CTX_BinOpR op e1) Γ TBool Γ Ï„ - | TP_CTX_BinOpL_PtrEq e2 Γ Ï„ : - typed Γ e2 (ref Ï„) → - typed_ctx_item (CTX_BinOpL EqOp e2) Γ (ref Ï„) Γ TBool - | TP_CTX_BinOpR_PtrEq e1 Γ Ï„ : - typed Γ e1 (ref Ï„) → - typed_ctx_item (CTX_BinOpR EqOp e1) Γ (ref Ï„) Γ TBool + | TP_CTX_BinOpL_UnboxedEq e2 Γ Ï„ : + UnboxedType Ï„ → + typed Γ e2 Ï„ → + typed_ctx_item (CTX_BinOpL EqOp e2) Γ Ï„ Γ TBool + | TP_CTX_BinOpR_UnboxedEq e1 Γ Ï„ : + UnboxedType Ï„ → + typed Γ e1 Ï„ → + typed_ctx_item (CTX_BinOpR EqOp e1) Γ Ï„ Γ TBool | TP_CTX_IfL Γ e1 e2 Ï„ : typed Γ e1 Ï„ → typed Γ e2 Ï„ → typed_ctx_item (CTX_IfL e1 e2) Γ (TBool) Γ Ï„ @@ -198,15 +200,15 @@ Inductive typed_ctx_item : Γ ⊢ₜ e1 : TRef TNat → typed_ctx_item (CTX_FAAR e1) Γ TNat Γ TNat | TP_CTX_CasL Γ e1 e2 Ï„ : - EqType Ï„ → UnboxedType Ï„ → + UnboxedType Ï„ → typed Γ e1 Ï„ → typed Γ e2 Ï„ → typed_ctx_item (CTX_CmpXchgL e1 e2) Γ (TRef Ï„) Γ (TProd Ï„ TBool) | TP_CTX_CasM Γ e0 e2 Ï„ : - EqType Ï„ → UnboxedType Ï„ → + UnboxedType Ï„ → typed Γ e0 (TRef Ï„) → typed Γ e2 Ï„ → typed_ctx_item (CTX_CmpXchgM e0 e2) Γ Ï„ Γ (TProd Ï„ TBool) | TP_CTX_CasR Γ e0 e1 Ï„ : - EqType Ï„ → UnboxedType Ï„ → + UnboxedType Ï„ → typed Γ e0 (TRef Ï„) → typed Γ e1 Ï„ → typed_ctx_item (CTX_CmpXchgR e0 e1) Γ Ï„ Γ (TProd Ï„ TBool) (* Polymorphic & recursive types *) @@ -238,7 +240,8 @@ Inductive typed_ctx: ctx → stringmap type → type → stringmap type → type typed_ctx K Γ1 Ï„1 Γ2 Ï„2 → typed_ctx (k :: K) Γ1 Ï„1 Γ3 Ï„3. -(* Observable types are, at the moment, exactly the types which support equality. *) +(* Observable types are, at the moment, exactly the unboxed types +which support direct equality tests. *) Definition ObsType : type → Prop := λ Ï„, EqType Ï„ ∧ UnboxedType Ï„. Definition ctx_refines (Γ : stringmap type) @@ -504,9 +507,9 @@ Section bin_log_related_under_typed_ctx. + iApply bin_log_related_bool_binop; try (by iApply fundamental); eauto. by iApply (IHK with "Hrel"). - + iApply bin_log_related_ref_binop; try (by iApply fundamental). + + iApply bin_log_related_unboxed_eq; try (eassumption || by iApply fundamental). by iApply (IHK with "Hrel"). - + iApply bin_log_related_ref_binop; try (by iApply fundamental). + + iApply bin_log_related_unboxed_eq; try (eassumption || by iApply fundamental). by iApply (IHK with "Hrel"). + iApply (bin_log_related_if with "[] []"); try by iApply fundamental. diff --git a/theories/typing/fundamental.v b/theories/typing/fundamental.v index 87eefa61e53e20548eced215c1c6b95868078335..c51d82cb53cc0b3e5b82706ea4118b602b426fed 100644 --- a/theories/typing/fundamental.v +++ b/theories/typing/fundamental.v @@ -288,7 +288,7 @@ Section fundamental. rel_values. Qed. - Lemma bin_log_related_CmpXchg Δ Γ e1 e2 e3 e1' e2' e3' Ï„ + Lemma bin_log_related_CmpXchg_EqType Δ Γ e1 e2 e3 e1' e2' e3' Ï„ (HEqÏ„ : EqType Ï„) (HUbÏ„ : UnboxedType Ï„) : ({Δ;Γ} ⊨ e1 ≤log≤ e1' : TRef Ï„) -∗ @@ -327,6 +327,24 @@ Section fundamental. iFrame "Hv". iExists _. done. Qed. + Lemma bin_log_related_CmpXchg Δ Γ e1 e2 e3 e1' e2' e3' Ï„ + (HUbÏ„ : UnboxedType Ï„) : + ({Δ;Γ} ⊨ e1 ≤log≤ e1' : TRef Ï„) -∗ + ({Δ;Γ} ⊨ e2 ≤log≤ e2' : Ï„) -∗ + ({Δ;Γ} ⊨ e3 ≤log≤ e3' : Ï„) -∗ + {Δ;Γ} ⊨ CmpXchg e1 e2 e3 ≤log≤ CmpXchg e1' e2' e3' : TProd Ï„ TBool. + Proof. + cut (EqType Ï„ ∨ ∃ Ï„', Ï„ = TRef Ï„'). + { intros [HÏ„ | [Ï„' ->]]. + - by iApply bin_log_related_CmpXchg_EqType. + - iIntros "H1 H2 H3". intro_clause. + iSpecialize ("H1" with "Hvs"). + iSpecialize ("H2" with "Hvs"). + iSpecialize ("H3" with "Hvs"). + iApply (refines_cmpxchg_ref with "H1 H2 H3"). } + by apply unboxed_type_ref_or_eqtype. + Qed. + Lemma bin_log_related_alloc Δ Γ e e' Ï„ : ({Δ;Γ} ⊨ e ≤log≤ e' : Ï„) -∗ {Δ;Γ} ⊨ Alloc e ≤log≤ Alloc e' : TRef Ï„. @@ -342,31 +360,25 @@ Section fundamental. rel_values. iExists l, k. eauto. Qed. - Lemma bin_log_related_ref_binop Δ Γ e1 e2 e1' e2' Ï„ : - ({Δ;Γ} ⊨ e1 ≤log≤ e1' : TRef Ï„) -∗ - ({Δ;Γ} ⊨ e2 ≤log≤ e2' : TRef Ï„) -∗ + Lemma bin_log_related_unboxed_eq Δ Γ e1 e2 e1' e2' Ï„ : + UnboxedType Ï„ → + ({Δ;Γ} ⊨ e1 ≤log≤ e1' : Ï„) -∗ + ({Δ;Γ} ⊨ e2 ≤log≤ e2' : Ï„) -∗ {Δ;Γ} ⊨ BinOp EqOp e1 e2 ≤log≤ BinOp EqOp e1' e2' : TBool. Proof. - iIntros "IH1 IH2". + iIntros (HÏ„) "IH1 IH2". intro_clause. rel_bind_ap e2 e2' "IH2" v2 v2' "#IH2". rel_bind_ap e1 e1' "IH1" v1 v1' "#IH1". - iDestruct "IH1" as (l1 l2) "[% [% #Hl]]"; simplify_eq/=. - iDestruct "IH2" as (l1' l2') "[% [% #Hl']]"; simplify_eq/=. + iAssert (⌜val_is_unboxed v1âŒ)%I as "%". + { rewrite !unboxed_type_sound //. + iDestruct "IH1" as "[$ _]". } + iAssert (⌜val_is_unboxed v2'âŒ)%I as "%". + { rewrite !unboxed_type_sound //. + iDestruct "IH2" as "[_ $]". } + iMod (unboxed_type_eq with "IH1 IH2") as "%"; first done. rel_op_l. rel_op_r. - destruct (decide (l1 = l1')) as [->|?]. - - iMod (interp_ref_funct _ _ l1' l2 l2' with "[Hl Hl']") as %->. - { solve_ndisj. } - { iSplit; iExists _,_; eauto. } - rewrite !bool_decide_true //. - value_case. - - destruct (decide (l2 = l2')) as [->|?]. - + iMod (interp_ref_inj _ _ l2' l1 l1' with "[Hl Hl']") as %->. - { solve_ndisj. } - { iSplit; iExists _,_; eauto. } - congruence. - + rewrite !bool_decide_false //; try congruence. - value_case. + do 2 case_bool_decide; first [by rel_values | naive_solver]. Qed. Lemma bin_log_related_nat_binop Δ Γ op e1 e2 e1' e2' Ï„ : @@ -531,7 +543,7 @@ Section fundamental. by iApply fundamental. + iApply bin_log_related_bool_unop; first done. by iApply fundamental. - + iApply bin_log_related_ref_binop; + + iApply bin_log_related_unboxed_eq; try done; by iApply fundamental. + iApply bin_log_related_pair; by iApply fundamental. diff --git a/theories/typing/interp.v b/theories/typing/interp.v index 3f6e230b48e8f31a45c724174cb5c5ad3add8dc1..403a9ed0653aa879e30ea6133ae2134627991dbb 100644 --- a/theories/typing/interp.v +++ b/theories/typing/interp.v @@ -74,6 +74,35 @@ Section semtypes. + iDestruct "H1" as "[% [% H1]]"; simplify_eq/=. rewrite IHHÏ„2. by iDestruct "H1" as "%"; subst. Qed. + + Lemma unboxed_type_eq Ï„ Δ v1 v2 w1 w2 : + UnboxedType Ï„ → + interp Ï„ Δ v1 v2 -∗ + interp Ï„ Δ w1 w2 -∗ + |={⊤}=> ⌜v1 = w1 ↔ v2 = w2âŒ. + Proof. + intros Hunboxed. + cut (EqType Ï„ ∨ ∃ Ï„', Ï„ = TRef Ï„'). + { intros [HÏ„ | [Ï„' ->]]. + - rewrite !eq_type_sound //. + iIntros "% %". iModIntro. + iPureIntro. naive_solver. + - rewrite /lrel_car /=. + iDestruct 1 as (l1 l2 -> ->) "Hl". + iDestruct 1 as (r1 r2 -> ->) "Hr". + destruct (decide (l1 = r1)); subst. + + destruct (decide (l2 = r2)); subst; first by eauto. + iInv (relocN.@"ref".@(r1, l2)) as (v1 v2) "(>Hr1 & >Hr2 & Hinv1)". + iInv (relocN.@"ref".@(r1, r2)) as (w1 w2) "(>Hr1' & >Hr2' & Hinv2)". + iExFalso. by iDestruct (mapsto_valid_2 with "Hr1 Hr1'") as %[]. + + destruct (decide (l2 = r2)); subst; last first. + { iModIntro. iPureIntro. naive_solver. } + iInv (relocN.@"ref".@(r1, r2)) as (v1 v2) "(>Hr1 & >Hr2 & Hinv1)". + iInv (relocN.@"ref".@(l1, r2)) as (w1 w2) "(>Hr1' & >Hr2' & Hinv2)". + iExFalso. by iDestruct (mapstoS_valid_2 with "Hr2 Hr2'") as %[]. } + by apply unboxed_type_ref_or_eqtype. + Qed. + End semtypes. (** ** Properties of the type inrpretation w.r.t. the substitutions *) diff --git a/theories/typing/soundness.v b/theories/typing/soundness.v index 0e821dd6b8ff1ce2709600ff2f593cd56db8d42a..150312044f9c1c8dec1ab6dfed9daffe9fd5cc0c 100644 --- a/theories/typing/soundness.v +++ b/theories/typing/soundness.v @@ -8,7 +8,7 @@ Lemma logrel_adequate Σ `{relocPreG Σ} e e' Ï„ (σ : state) : (∀ `{relocG Σ} Δ, ⊢ {⊤;Δ;∅} ⊨ e ≤log≤ e' : Ï„) → adequate NotStuck e σ (λ v _, ∃ thp' h v', rtc erased_step ([e'], σ) (of_val v' :: thp', h) - ∧ (ObsType Ï„ → v = v')). + ∧ (EqType Ï„ → v = v')). Proof. intros Hlog. set (A := λ (HΣ : relocG Σ), interp Ï„ []). @@ -20,7 +20,6 @@ Proof. { iApply env_ltyped2_empty. } by rewrite !fmap_empty !subst_map_empty. - intros HΣ v v'. unfold A. iIntros "Hvv". - unfold ObsType. cbn. iIntros (HÏ„). by iApply (eq_type_sound with "Hvv"). Qed. @@ -30,7 +29,7 @@ Theorem logrel_typesafety Σ `{relocPreG Σ} e e' Ï„ thp σ σ' : not_stuck e' σ'. Proof. intros Hlog ??. - cut (adequate NotStuck e σ (λ v _, ∃ thp' h v', rtc erased_step ([e], σ) (of_val v' :: thp', h) ∧ (ObsType Ï„ → v = v'))); first (intros [_ ?]; eauto). + cut (adequate NotStuck e σ (λ v _, ∃ thp' h v', rtc erased_step ([e], σ) (of_val v' :: thp', h) ∧ (EqType Ï„ → v = v'))); first (intros [_ ?]; eauto). eapply logrel_adequate; eauto. Qed. @@ -51,8 +50,8 @@ Lemma logrel_simul Σ `{relocPreG Σ} (∃ thp' hp' v', rtc erased_step ([e'], σ) (of_val v' :: thp', hp') ∧ (ObsType Ï„ → v = v')). Proof. intros Hlog Hsteps. - cut (adequate NotStuck e σ (λ v _, ∃ thp' h v', rtc erased_step ([e'], σ) (of_val v' :: thp', h) ∧ (ObsType Ï„ → v = v'))). - { destruct 1; naive_solver. } + cut (adequate NotStuck e σ (λ v _, ∃ thp' h v', rtc erased_step ([e'], σ) (of_val v' :: thp', h) ∧ (EqType Ï„ → v = v'))). + { unfold ObsType. destruct 1; naive_solver. } eapply logrel_adequate; eauto. Qed. diff --git a/theories/typing/types.v b/theories/typing/types.v index 574390092bd740874387119b6c55a14a3140bf3b..678eb6c8d373cc1baa3e0dd232a18f092677df5f 100644 --- a/theories/typing/types.v +++ b/theories/typing/types.v @@ -18,7 +18,15 @@ Inductive type := | TExists : {bind 1 of type} → type | TRef : type → type. -(** Types which support direct equality test (which coincides with ctx equiv) *) +(** Which types are unboxed -- we can only do CAS on locations which hold unboxed types *) +Inductive UnboxedType : type → Prop := + | UnboxedTUnit : UnboxedType TUnit + | UnboxedTNat : UnboxedType TNat + | UnboxedTBool : UnboxedType TBool + | UnboxedTRef Ï„ : UnboxedType (TRef Ï„). + +(** Types which support direct equality test (which coincides with ctx equiv). + This is an auxiliary notion. *) Inductive EqType : type → Prop := | EqTUnit : EqType TUnit | EqTNat : EqType TNat @@ -26,12 +34,9 @@ Inductive EqType : type → Prop := | EqTProd Ï„ Ï„' : EqType Ï„ → EqType Ï„' → EqType (TProd Ï„ Ï„') | EqSum Ï„ Ï„' : EqType Ï„ → EqType Ï„' → EqType (TSum Ï„ Ï„'). -(** Which types are unboxed -- we can only do CAS on locations which hold unboxed types *) -Inductive UnboxedType : type → Prop := - | UnboxedTUnit : UnboxedType TUnit - | UnboxedTNat : UnboxedType TNat - | UnboxedTBool : UnboxedType TBool - | UnboxedTRef Ï„ : UnboxedType (TRef Ï„). +Lemma unboxed_type_ref_or_eqtype Ï„ : + UnboxedType Ï„ → (EqType Ï„ ∨ ∃ Ï„', Ï„ = TRef Ï„'). +Proof. inversion 1; first [ left; by econstructor | right; naive_solver ]. Qed. (** Autosubst instances *) Instance Ids_type : Ids type. derive. Defined. @@ -135,8 +140,9 @@ Inductive typed : stringmap type → expr → type → Prop := Γ ⊢ₜ e : TBool → unop_bool_res_type op = Some Ï„ → Γ ⊢ₜ UnOp op e : Ï„ - | RefEq_typed Γ e1 e2 Ï„ : - Γ ⊢ₜ e1 : ref Ï„ → Γ ⊢ₜ e2 : ref Ï„ → + | UnboxedEq_typed Γ e1 e2 Ï„ : + UnboxedType Ï„ → + Γ ⊢ₜ e1 : Ï„ → Γ ⊢ₜ e2 : Ï„ → Γ ⊢ₜ BinOp EqOp e1 e2 : TBool | Pair_typed Γ e1 e2 Ï„1 Ï„2 : Γ ⊢ₜ e1 : Ï„1 → Γ ⊢ₜ e2 : Ï„2 → @@ -198,7 +204,7 @@ Inductive typed : stringmap type → expr → type → Prop := Γ ⊢ₜ e2 : TNat → Γ ⊢ₜ FAA e1 e2 : TNat | TCmpXchg Γ e1 e2 e3 Ï„ : - EqType Ï„ → UnboxedType Ï„ → + UnboxedType Ï„ → Γ ⊢ₜ e1 : ref Ï„ → Γ ⊢ₜ e2 : Ï„ → Γ ⊢ₜ e3 : Ï„ → Γ ⊢ₜ CmpXchg e1 e2 e3 : Ï„ * TBool with val_typed : val → type → Prop :=