diff --git a/iris_heap_lang/primitive_laws.v b/iris_heap_lang/primitive_laws.v index 9b3510038a68406d9a226c5d82c8ab18a1c072ce..649c64383ab20ce935b7cb895e45af269d68c02a 100644 --- a/iris_heap_lang/primitive_laws.v +++ b/iris_heap_lang/primitive_laws.v @@ -317,7 +317,7 @@ Lemma twp_xchg s E l v' v : [[{ RET v'; l ↦ v }]]. Proof. iIntros (Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; first done. - iIntros (σ1 κs n) "[Hσ Hκs] !>". iDestruct (gen_heap_valid with "Hσ Hl") as %?. + iIntros (σ1 ns κs nt) "[Hσ Hκs] !>". iDestruct (gen_heap_valid with "Hσ Hl") as %?. iSplit; first by eauto with head_step. iIntros (κ v2 σ2 efs Hstep); inv_head_step. iMod (gen_heap_update with "Hσ Hl") as "[$ Hl]". diff --git a/iris_staging/heap_lang/interpreter.v b/iris_staging/heap_lang/interpreter.v index e0f8f1c17ea1430f31297843a9084de78e79ffc7..8d7e31a97105ed57ba50d9fdef6034f27d46a8b9 100644 --- a/iris_staging/heap_lang/interpreter.v +++ b/iris_staging/heap_lang/interpreter.v @@ -535,6 +535,13 @@ Section interpreter. let '(l, _) := l_v0 in _ ↠interp_modify_state (state_upd_heap <[l:=Some w]>); mret (LitV LitUnit) + | Xchg el e => + w ↠interp e; + vl ↠interp el; + l_v0 ↠read_loc "xchg" vl; + let '(l, v0) := l_v0 in + _ ↠interp_modify_state (state_upd_heap <[l:=Some w]>); + mret v0 | CmpXchg e e1 e2 => v2 ↠interp e2; v1 ↠interp e1; diff --git a/tests/heap_lang.ref b/tests/heap_lang.ref index e61cca8e1e6d475d17e33bda8a885296814b785d..64e37382064e716331fc28ad29f4b1d9f4bea613 100644 --- a/tests/heap_lang.ref +++ b/tests/heap_lang.ref @@ -267,37 +267,6 @@ goal 2 is: --------------------------------------∗ WP #() #() @ E ∖ ↑N {{ _, |={E ∖ ↑N}=> â–· P ∗ True }} -"xchg_example" - : string -1 goal - - Σ : gFunctors - heapG0 : heapG Σ - l : loc - vâ‚, vâ‚‚ : val - Φ : val → iPropI Σ - ============================ - "Hl" : l ↦ vâ‚‚ - "HΦ" : l ↦ vâ‚‚ -∗ Φ vâ‚ - --------------------------------------∗ - |={⊤}=> Φ vâ‚ - -1 goal - - Σ : gFunctors - heapG0 : heapG Σ - N : namespace - l : loc - v : val - Φ : language.val heap_lang → iPropI Σ - v' : val - ============================ - "HΦ" : ∀ v'0 : language.val heap_lang, True -∗ Φ v'0 - "Hl" : l ↦ v - "Hclose" : â–· (∃ v0 : val, l ↦ v0) ={⊤ ∖ ↑N,⊤}=∗ emp - --------------------------------------∗ - |={⊤ ∖ ↑N,⊤}=> Φ v' - 1 goal Σ : gFunctors diff --git a/tests/heap_lang.v b/tests/heap_lang.v index 458755adedb4bfa33a7e31e5d0d38f94793ad1c8..cb7c84810285a63b361ba9d7ef6d8639528a8cb1 100644 --- a/tests/heap_lang.v +++ b/tests/heap_lang.v @@ -199,6 +199,37 @@ Section tests. iIntros (?) "?". wp_cmpxchg as ? | ?. done. done. Qed. + Lemma wp_xchg l (vâ‚ vâ‚‚ : val) : + {{{ l ↦ vâ‚ }}} + Xchg #l vâ‚‚ + {{{ RET vâ‚; l ↦ vâ‚‚ }}}. + Proof. + iIntros (Φ) "Hl HΦ". + wp_xchg. + iApply "HΦ" => //. + Qed. + + Lemma twp_xchg l (vâ‚ vâ‚‚ : val) : + l ↦ vâ‚ -∗ + WP Xchg #l vâ‚‚ [{ vâ‚, l ↦ vâ‚‚ }]. + Proof. + iIntros "Hl". + wp_xchg => //. + Qed. + + Lemma wp_xchg_inv N l (v : val) : + {{{ inv N (∃ v, l ↦ v) }}} + Xchg #l v + {{{ v', RET v'; True }}}. + Proof. + iIntros (Φ) "Hl HΦ". + iInv "Hl" as (v') "Hl" "Hclose". + wp_xchg. + iApply "HΦ". + iApply "Hclose". + iExists _ => //. + Qed. + Lemma wp_alloc_split : ⊢ WP Alloc #0 {{ _, True }}. Proof. wp_alloc l as "[Hl1 Hl2]". Show. done. Qed. @@ -383,29 +414,6 @@ Section atomic. iIntros (?) "H". iInv "H" as "H". Show. Abort. - Check "xchg_example". - - Lemma xchg_example l (vâ‚ vâ‚‚ : val) : - {{{ l ↦ vâ‚ }}} - Xchg #l vâ‚‚ - {{{ RET vâ‚; l ↦ vâ‚‚ }}}. - Proof. - iIntros (Φ) "Hl HΦ". - wp_xchg. - Show. - Abort. - - Lemma xchg_inv_example N l (v : val) : - {{{ inv N (∃ v, l ↦ v) }}} - Xchg #l v - {{{ v', RET v'; True }}}. - Proof. - iIntros (Φ) "Hl HΦ". - iInv "Hl" as (v') "Hl" "Hclose". - wp_xchg. - Show. - Abort. - End atomic. Section printing_tests.