diff --git a/theories/base_logic/lib/invariants.v b/theories/base_logic/lib/invariants.v index 1ecc9bcc72d82703cc2feaa1231ecf97b684c1df..1405eaf7860b0a2c38ac52b3d33f940f85427139 100644 --- a/theories/base_logic/lib/invariants.v +++ b/theories/base_logic/lib/invariants.v @@ -102,8 +102,8 @@ Proof. rewrite difference_diag_L. iPoseProof (fupd_mask_frame_r _ _ (E ∖ ↑ N) with "H") as "H"; first set_solver. rewrite left_id_L -union_difference_L //. iMod "H" as "[$ H]"; iModIntro. - iIntros (E') "HP". iSpecialize ("H" with "HP"). - iPoseProof (fupd_mask_frame_r _ _ E' with "H") as "H"; first set_solver. + iIntros (E') "HP". + iPoseProof (fupd_mask_frame_r _ _ E' with "(H HP)") as "H"; first set_solver. by rewrite left_id_L. Qed. diff --git a/theories/heap_lang/lib/par.v b/theories/heap_lang/lib/par.v index 47498b0d04813dc800988e0461be5d31be23c7b8..5fd18ec3f008159b067604eca73d700595c6bf18 100644 --- a/theories/heap_lang/lib/par.v +++ b/theories/heap_lang/lib/par.v @@ -31,7 +31,7 @@ Proof. iIntros (l) "Hl". wp_let. wp_bind (f2 _). wp_apply (wp_wand with "Hf2"); iIntros (v) "H2". wp_let. wp_apply (join_spec with "[$Hl]"). iIntros (w) "H1". - iSpecialize ("HΦ" with "[-]"); first by iSplitL "H1". by wp_pures. + iSpecialize ("HΦ" with "[$H1 $H2]"). by wp_pures. Qed. Lemma wp_par (Ψ1 Ψ2 : val → iProp Σ) (e1 e2 : expr) (Φ : val → iProp Σ) : diff --git a/theories/program_logic/adequacy.v b/theories/program_logic/adequacy.v index fdbb25a23294b0708b2f1a72c675da22772141c2..770e56fb1b01dbd4ef41864d71a7d772e5174947 100644 --- a/theories/program_logic/adequacy.v +++ b/theories/program_logic/adequacy.v @@ -168,8 +168,7 @@ Proof. iIntros (?) "Hφ Hσ He Ht". rewrite Nat_iter_S_r. iDestruct (wptp_steps _ n with "Hσ He Ht") as "H"; first done. iApply (step_fupdN_wand with "H"). iDestruct 1 as (e2' t2' ->) "[Hσ _]". - iSpecialize ("Hφ" with "Hσ"). - iMod (fupd_plain_mask_empty _ ⌜φâŒ%I with "Hφ") as %?. + iMod (fupd_plain_mask_empty _ ⌜φâŒ%I with "(Hφ Hσ)") as %?. by iApply step_fupd_intro. Qed. End adequacy. diff --git a/theories/program_logic/hoare.v b/theories/program_logic/hoare.v index f75756fa6f60781987fe52f4739eb84d839c7125..844842a897de86773c3b0ff894b02abfb6108ef6 100644 --- a/theories/program_logic/hoare.v +++ b/theories/program_logic/hoare.v @@ -75,7 +75,7 @@ Lemma ht_vs s E P P' Φ Φ' e : ⊢ {{ P }} e @ s; E {{ Φ }}. Proof. iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iMod ("Hvs" with "HP") as "HP". - iApply wp_fupd. iApply (wp_wand with "[HP]"); [by iApply "Hwp"|]. + iApply wp_fupd. iApply (wp_wand with "(Hwp HP)"). iIntros (v) "Hv". by iApply "HΦ". Qed. @@ -85,7 +85,7 @@ Lemma ht_atomic s E1 E2 P P' Φ Φ' e `{!Atomic (stuckness_to_atomicity s) e} : Proof. iIntros "(#Hvs & #Hwp & #HΦ) !# HP". iApply (wp_atomic _ _ E2); auto. iMod ("Hvs" with "HP") as "HP". iModIntro. - iApply (wp_wand with "[HP]"); [by iApply "Hwp"|]. + iApply (wp_wand with "(Hwp HP)"). iIntros (v) "Hv". by iApply "HΦ". Qed. @@ -94,7 +94,7 @@ Lemma ht_bind `{LanguageCtx Λ K} s E P Φ Φ' e : ⊢ {{ P }} K e @ s; E {{ Φ' }}. Proof. iIntros "[#Hwpe #HwpK] !# HP". iApply wp_bind. - iApply (wp_wand with "[HP]"); [by iApply "Hwpe"|]. + iApply (wp_wand with "(Hwpe HP)"). iIntros (v) "Hv". by iApply "HwpK". Qed. diff --git a/theories/program_logic/total_adequacy.v b/theories/program_logic/total_adequacy.v index c742a636dc8e7c288049abf01b5de88b3a77d2bc..b6243370d78081ed8aa584f9b67f0d0606331758 100644 --- a/theories/program_logic/total_adequacy.v +++ b/theories/program_logic/total_adequacy.v @@ -94,12 +94,12 @@ Proof. iMod ("IH" with "Hσ1") as "[_ IH]". iMod ("IH" with "[% //]") as "($ & Hσ & [IH _] & IHfork)". iModIntro. iExists (length efs + n). iFrame "Hσ". - iApply (twptp_app [_] with "[IH]"); first by iApply "IH". + iApply (twptp_app [_] with "(IH [//])"). clear. iInduction efs as [|e efs] "IH"; simpl. { rewrite twptp_unfold /twptp_pre. iIntros (t2 σ1 κ κs σ2 n1 Hstep). destruct Hstep; simplify_eq/=; discriminate_list. } iDestruct "IHfork" as "[[IH' _] IHfork]". - iApply (twptp_app [_] with "[IH']"). by iApply "IH'". by iApply "IH". + iApply (twptp_app [_] with "(IH' [//])"). by iApply "IH". Qed. Lemma twptp_total n σ t : diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v index 8cfeea29fd756affef44c8502ca562afa8cdf345..92ddffd1616eac02a9e906c36d9fbe860695fdbe 100644 --- a/theories/program_logic/weakestpre.v +++ b/theories/program_logic/weakestpre.v @@ -292,7 +292,7 @@ Section proofmode_classes. Proof. intros ?. rewrite /ElimAcc. iIntros "Hinner >Hacc". iDestruct "Hacc" as (x) "[Hα Hclose]". - iApply (wp_wand with "[Hinner Hα]"); first by iApply "Hinner". + iApply (wp_wand with "(Hinner Hα)"). iIntros (v) ">[Hβ HΦ]". iApply "HΦ". by iApply "Hclose". Qed. @@ -304,7 +304,7 @@ Section proofmode_classes. rewrite /ElimAcc. iIntros "Hinner >Hacc". iDestruct "Hacc" as (x) "[Hα Hclose]". iApply wp_fupd. - iApply (wp_wand with "[Hinner Hα]"); first by iApply "Hinner". + iApply (wp_wand with "(Hinner Hα)"). iIntros (v) ">[Hβ HΦ]". iApply "HΦ". by iApply "Hclose". Qed. End proofmode_classes.