diff --git a/tests/heap_lang.v b/tests/heap_lang.v index 4210930212e97c29808c1daa9edeb63f03982fdd..42940fee7b9b09305406e32eb07482dee1edabb0 100644 --- a/tests/heap_lang.v +++ b/tests/heap_lang.v @@ -142,6 +142,10 @@ Section tests. P -∗ (∀ Q Φ, Q -∗ WP e {{ Φ }}) -∗ WP e {{ _, True }}. Proof. iIntros "HP HW". wp_apply "HW". iExact "HP". Qed. + Lemma wp_pures_val (b : bool) : + ⊢ WP #b {{ _, True }}. + Proof. wp_pures. done. Qed. + Lemma wp_cmpxchg l v : val_is_unboxed v → l ↦ v -∗ WP CmpXchg #l v v {{ _, True }}. diff --git a/theories/heap_lang/proofmode.v b/theories/heap_lang/proofmode.v index d24214eb7d8e8d249014237e7ee17df51914de6c..b5a5f78ccbb117e334d9640a2bbd13fa68c820ed 100644 --- a/theories/heap_lang/proofmode.v +++ b/theories/heap_lang/proofmode.v @@ -107,8 +107,10 @@ Tactic Notation "wp_pure" open_constr(efoc) := (* TODO: do this in one go, without [repeat]. *) Ltac wp_pures := iStartProof; - repeat (wp_pure _; []). (* The `;[]` makes sure that no side-condition - magically spawns. *) + first [ (* The `;[]` makes sure that no side-condition magically spawns. *) + progress repeat (wp_pure _; []) + | wp_finish (* In case wp_pure never ran, make sure we do the usual cleanup. *) + ]. (** Unlike [wp_pures], the tactics [wp_rec] and [wp_lam] should also reduce lambdas/recs that are hidden behind a definition, i.e. they should use