diff --git a/theories/heap_lang/lang.v b/theories/heap_lang/lang.v index 538eafe00673a60daaa0e24496fda826215030cb..4c0027c96c5afc8a3ebc2f3a93eaa58252444616 100644 --- a/theories/heap_lang/lang.v +++ b/theories/heap_lang/lang.v @@ -361,6 +361,23 @@ Definition bin_op_eval (op : bin_op) (v1 v2 : val) : option val := | _, _ => None end. +(** Return whether it is possible to use CAS to compare vl (current value) with v1 (netest value). *) +Definition vals_cas_compare_safe (vl v1 : val) : Prop := + match vl, v1 with + | LitV _, LitV _ => True + (* We want to support CAS'ing [NONEV := InjLV LitUnit] to [SOMEV #l := InjRV + (LitV l)]. An implementation of this is possible if literals have an invalid + bit pattern that can be used to represent NONE. *) + | InjLV (LitV LitUnit), InjLV (LitV LitUnit) => True + | InjLV (LitV LitUnit), InjRV (LitV _) => True + | InjRV (LitV _), InjLV (LitV LitUnit) => True + | _, _ => False + end. +(** Just a sanity check. *) +Lemma vals_cas_compare_safe_sym vl v1 : + vals_cas_compare_safe vl v1 → vals_cas_compare_safe v1 vl. +Proof. rewrite /vals_cas_compare_safe. repeat case_match; done. Qed. + Inductive head_step : expr → state → expr → state → list (expr) → Prop := | BetaS f x e1 e2 v2 e' σ : to_val e2 = Some v2 → @@ -405,10 +422,12 @@ Inductive head_step : expr → state → expr → state → list (expr) → Prop | CasFailS l e1 v1 e2 v2 vl σ : to_val e1 = Some v1 → to_val e2 = Some v2 → σ !! l = Some vl → vl ≠v1 → + vals_cas_compare_safe vl v1 → head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool false) σ [] | CasSucS l e1 v1 e2 v2 σ : to_val e1 = Some v1 → to_val e2 = Some v2 → σ !! l = Some v1 → + vals_cas_compare_safe v1 v1 → head_step (CAS (Lit $ LitLoc l) e1 e2) σ (Lit $ LitBool true) (<[l:=v2]>σ) [] | FaaS l i1 e2 i2 σ : to_val e2 = Some (LitV (LitInt i2)) → diff --git a/theories/heap_lang/lib/atomic_heap.v b/theories/heap_lang/lib/atomic_heap.v index bc4160c9b620dcf3bd952ad6fa490e4f6e781f0f..9de110430c2391ff9789cd62b8731dd50c72a9c0 100644 --- a/theories/heap_lang/lib/atomic_heap.v +++ b/theories/heap_lang/lib/atomic_heap.v @@ -37,7 +37,7 @@ Structure atomic_heap {Σ} `{!heapG Σ} := AtomicHeap { IntoVal e1 w1 → IntoVal e2 w2 → atomic_wp (cas (#l, e1, e2))%E ⊤ ⊤ - (λ v, mapsto l 1 v) + (λ v, ⌜vals_cas_compare_safe v w1⌠∗ mapsto l 1 v)%I (λ v (_:()), if decide (v = w1) then mapsto l 1 w2 else mapsto l 1 v) (λ v _, #(if decide (v = w1) then true else false)%V); }. @@ -91,12 +91,12 @@ Section proof. IntoVal e1 w1 → IntoVal e2 w2 → atomic_wp (primitive_cas (#l, e1, e2))%E ⊤ ⊤ - (λ v, l ↦ v)%I + (λ v, ⌜vals_cas_compare_safe v w1⌠∗ l ↦ v)%I (λ v (_:()), if decide (v = w1) then l ↦ w2 else l ↦ v)%I (λ v _, #(if decide (v = w1) then true else false)%V). Proof. iIntros (<- <- Q Φ) "? AU". wp_let. repeat wp_proj. - iMod (aupd_acc with "AU") as (v) "[H↦ [_ Hclose]]"; first solve_ndisj. + iMod (aupd_acc with "AU") as (v) "[[% H↦] [_ Hclose]]"; first solve_ndisj. destruct (decide (v = w1)) as [Hv|Hv]; [wp_cas_suc|wp_cas_fail]; iMod ("Hclose" $! () with "H↦") as "HΦ"; by iApply "HΦ". Qed. diff --git a/theories/heap_lang/lib/increment.v b/theories/heap_lang/lib/increment.v index 938b2d6b20d7628e5de2652f52d6c56f4bca05a2..d2db12190118428c6ce2cdb6dde671a1b563086f 100644 --- a/theories/heap_lang/lib/increment.v +++ b/theories/heap_lang/lib/increment.v @@ -39,8 +39,8 @@ Section increment. (* Prove the atomic shift for CAS *) iAuIntro. iApply (aacc_aupd with "AU"); first done. iIntros (x') "H↦". - iApply (aacc_intro with "[H↦]"); [solve_ndisj|done|iSplit]. - { iIntros "$ !> $ !> //". } + iApply (aacc_intro with "[H↦]"); [solve_ndisj|simpl;by auto with iFrame|iSplit]. + { iIntros "[_ $] !> $ !> //". } iIntros ([]) "H↦ !>". destruct (decide (#x' = #x)) as [[= ->]|Hx]. - iRight. iExists (). iFrame. iIntros "HΦ !> HQ". diff --git a/theories/heap_lang/lifting.v b/theories/heap_lang/lifting.v index 3822ee5220c9814ed425b57e111839dbee012f49..b6734455cd0c135483a48cd801de2f90891136ef 100644 --- a/theories/heap_lang/lifting.v +++ b/theories/heap_lang/lifting.v @@ -185,22 +185,22 @@ Proof. Qed. Lemma wp_cas_fail s E l q v' e1 v1 e2 : - IntoVal e1 v1 → AsVal e2 → v' ≠v1 → + IntoVal e1 v1 → AsVal e2 → v' ≠v1 → vals_cas_compare_safe v' v1 → {{{ ▷ l ↦{q} v' }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ RET LitV (LitBool false); l ↦{q} v' }}}. Proof. - iIntros (<- [v2 <-] ? Φ) ">Hl HΦ". + iIntros (<- [v2 <-] ?? Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. iModIntro; iSplit=> //. iFrame. by iApply "HΦ". Qed. Lemma twp_cas_fail s E l q v' e1 v1 e2 : - IntoVal e1 v1 → AsVal e2 → v' ≠v1 → + IntoVal e1 v1 → AsVal e2 → v' ≠v1 → vals_cas_compare_safe v' v1 → [[{ l ↦{q} v' }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ RET LitV (LitBool false); l ↦{q} v' }]]. Proof. - iIntros (<- [v2 <-] ? Φ) "Hl HΦ". + iIntros (<- [v2 <-] ?? Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -208,11 +208,11 @@ Proof. Qed. Lemma wp_cas_suc s E l e1 v1 e2 v2 : - IntoVal e1 v1 → IntoVal e2 v2 → + IntoVal e1 v1 → IntoVal e2 v2 → vals_cas_compare_safe v1 v1 → {{{ ▷ l ↦ v1 }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ RET LitV (LitBool true); l ↦ v2 }}}. Proof. - iIntros (<- <- Φ) ">Hl HΦ". + iIntros (<- <- ? Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. @@ -220,11 +220,11 @@ Proof. iModIntro. iSplit=>//. by iApply "HΦ". Qed. Lemma twp_cas_suc s E l e1 v1 e2 v2 : - IntoVal e1 v1 → IntoVal e2 v2 → + IntoVal e1 v1 → IntoVal e2 v2 → vals_cas_compare_safe v1 v1 → [[{ l ↦ v1 }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ RET LitV (LitBool true); l ↦ v2 }]]. Proof. - iIntros (<- <- Φ) "Hl HΦ". + iIntros (<- <- ? Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (σ1) "Hσ !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto. iIntros (v2' σ2 efs Hstep); inv_head_step. diff --git a/theories/heap_lang/proofmode.v b/theories/heap_lang/proofmode.v index 66694a32bce54ce966a3a46cd879f26c01781380..c46eafa574440188ab96a9dbc6bab6d08e68764f 100644 --- a/theories/heap_lang/proofmode.v +++ b/theories/heap_lang/proofmode.v @@ -224,18 +224,18 @@ Qed. Lemma tac_wp_cas_fail Δ Δ' s E i K l q v e1 v1 e2 Φ : IntoVal e1 v1 → AsVal e2 → MaybeIntoLaterNEnvs 1 Δ Δ' → - envs_lookup i Δ' = Some (false, l ↦{q} v)%I → v ≠v1 → + envs_lookup i Δ' = Some (false, l ↦{q} v)%I → v ≠v1 → vals_cas_compare_safe v v1 → envs_entails Δ' (WP fill K (Lit (LitBool false)) @ s; E {{ Φ }}) → envs_entails Δ (WP fill K (CAS (Lit (LitLoc l)) e1 e2) @ s; E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ??????. + rewrite envs_entails_eq=> ???????. rewrite -wp_bind. eapply wand_apply; first exact: wp_cas_fail. rewrite into_laterN_env_sound -later_sep envs_lookup_split //; simpl. by apply later_mono, sep_mono_r, wand_mono. Qed. Lemma tac_twp_cas_fail Δ s E i K l q v e1 v1 e2 Φ : IntoVal e1 v1 → AsVal e2 → - envs_lookup i Δ = Some (false, l ↦{q} v)%I → v ≠v1 → + envs_lookup i Δ = Some (false, l ↦{q} v)%I → v ≠v1 → vals_cas_compare_safe v v1 → envs_entails Δ (WP fill K (Lit (LitBool false)) @ s; E [{ Φ }]) → envs_entails Δ (WP fill K (CAS (Lit (LitLoc l)) e1 e2) @ s; E [{ Φ }]). Proof. @@ -247,19 +247,19 @@ Qed. Lemma tac_wp_cas_suc Δ Δ' Δ'' s E i K l v e1 v1 e2 v2 Φ : IntoVal e1 v1 → IntoVal e2 v2 → MaybeIntoLaterNEnvs 1 Δ Δ' → - envs_lookup i Δ' = Some (false, l ↦ v)%I → v = v1 → + envs_lookup i Δ' = Some (false, l ↦ v)%I → v = v1 → vals_cas_compare_safe v v1 → envs_simple_replace i false (Esnoc Enil i (l ↦ v2)) Δ' = Some Δ'' → envs_entails Δ'' (WP fill K (Lit (LitBool true)) @ s; E {{ Φ }}) → envs_entails Δ (WP fill K (CAS (Lit (LitLoc l)) e1 e2) @ s; E {{ Φ }}). Proof. - rewrite envs_entails_eq=> ???????; subst. + rewrite envs_entails_eq=> ????????; subst. rewrite -wp_bind. eapply wand_apply; first exact: wp_cas_suc. rewrite into_laterN_env_sound -later_sep envs_simple_replace_sound //; simpl. rewrite right_id. by apply later_mono, sep_mono_r, wand_mono. Qed. Lemma tac_twp_cas_suc Δ Δ' s E i K l v e1 v1 e2 v2 Φ : IntoVal e1 v1 → IntoVal e2 v2 → - envs_lookup i Δ = Some (false, l ↦ v)%I → v = v1 → + envs_lookup i Δ = Some (false, l ↦ v)%I → v = v1 → vals_cas_compare_safe v v1 → envs_simple_replace i false (Esnoc Enil i (l ↦ v2)) Δ = Some Δ' → envs_entails Δ' (WP fill K (Lit (LitBool true)) @ s; E [{ Φ }]) → envs_entails Δ (WP fill K (CAS (Lit (LitLoc l)) e1 e2) @ s; E [{ Φ }]). @@ -406,6 +406,7 @@ Tactic Notation "wp_cas_fail" := [iSolveTC |solve_mapsto () |try congruence + |done |simpl; try wp_value_head] | |- envs_entails _ (twp ?s ?E ?e ?Q) => first @@ -414,6 +415,7 @@ Tactic Notation "wp_cas_fail" := |fail 1 "wp_cas_fail: cannot find 'CAS' in" e]; [solve_mapsto () |try congruence + |done |wp_expr_simpl; try wp_value_head] | _ => fail "wp_cas_fail: not a 'wp'" end. @@ -432,6 +434,7 @@ Tactic Notation "wp_cas_suc" := [iSolveTC |solve_mapsto () |try congruence + |done |pm_reflexivity |simpl; try wp_value_head] | |- envs_entails _ (twp ?E ?e ?Q) => @@ -441,6 +444,7 @@ Tactic Notation "wp_cas_suc" := |fail 1 "wp_cas_suc: cannot find 'CAS' in" e]; [solve_mapsto () |try congruence + |done |pm_reflexivity |wp_expr_simpl; try wp_value_head] | _ => fail "wp_cas_suc: not a 'wp'"