diff --git a/proofmode/pviewshifts.v b/proofmode/pviewshifts.v index 013403208e937224eedac6686d31d9359d961f67..20de4de2e2a8fd9684c974273d1e494b5462e4dc 100644 --- a/proofmode/pviewshifts.v +++ b/proofmode/pviewshifts.v @@ -181,3 +181,6 @@ Tactic Notation "iTimeless" constr(H) := Tactic Notation "iTimeless" constr(H) "as" constr(Hs) := iTimeless H; iDestruct H as Hs. + +Hint Extern 2 (of_envs _ ⊢ _) => + match goal with |- _ ⊢ (|={_}=> _)%I => iPvsIntro end. diff --git a/tests/one_shot.v b/tests/one_shot.v index 31b1b2bbd44433b140818d0a2944782d8c94c4a9..768662ea387d2fb1cfdefd31c4ee3ec49d0f5254 100644 --- a/tests/one_shot.v +++ b/tests/one_shot.v @@ -52,7 +52,7 @@ Proof. iPvsIntro. iApply "Hf"; iSplit. - iIntros {n} "!". wp_let. iInv> N as "[[Hl Hγ]|H]"; last iDestruct "H" as {m} "[Hl Hγ]". - + iApply wp_pvs. wp_cas_suc. iSplitL; [|by iLeft; iPvsIntro]. + + iApply wp_pvs. wp_cas_suc. iSplitL; [|by iLeft]. iPvs (own_update with "Hγ") as "Hγ". { by apply cmra_update_exclusive with (y:=Cinr (DecAgree n)). } iPvsIntro; iRight; iExists n; by iSplitL "Hl". @@ -79,7 +79,7 @@ Proof. iCombine "Hγ" "Hγ'" as "Hγ". iDestruct (own_valid with "#Hγ") as %[=->]%dec_agree_op_inv. iSplitL "Hl"; [iRight; by eauto|]. - wp_match. iApply wp_assert'. wp_op=>?; iPvsIntro; simplify_eq/=; eauto. + wp_match. iApply wp_assert'. wp_op=>?; simplify_eq/=; eauto. Qed. Lemma hoare_one_shot (Φ : val → iProp) :