diff --git a/ProofMode.md b/ProofMode.md index 3e3f9c429863c8509a3faf057c5c81a2cb103330..17ccf47b877484647d486caa271fbe354b6d6069 100644 --- a/ProofMode.md +++ b/ProofMode.md @@ -251,7 +251,6 @@ _specification patterns_ to express splitting of hypotheses: `P`, as well the remaining goal. - `[%]` : This pattern can be used when eliminating `P -★ Q` when `P` is pure. It will generate a Coq goal for `P` and does not consume any hypotheses. -- `*` : instantiate all top-level universal quantifiers with meta variables. For example, given: diff --git a/theories/heap_lang/lib/par.v b/theories/heap_lang/lib/par.v index 95bc308be7d4ddbaada90f28b62ea5541d350244..0d125e5f709fe261811a0159e69d8a877d2d1467 100644 --- a/theories/heap_lang/lib/par.v +++ b/theories/heap_lang/lib/par.v @@ -33,7 +33,7 @@ Proof. iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _). iApply (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_let. + iSpecialize ("HΦ" with "[-]"); first by iSplitL "H1". by wp_let. Qed. Lemma wp_par (Ψ1 Ψ2 : val → iProp Σ) diff --git a/theories/program_logic/adequacy.v b/theories/program_logic/adequacy.v index ee330e7d508121c493e454c9d4dea0d620ba9d6b..2c9ce1d1eb1a2817bd0181cb402b3d5cbef7cb27 100644 --- a/theories/program_logic/adequacy.v +++ b/theories/program_logic/adequacy.v @@ -134,7 +134,7 @@ Lemma wp_safe e σ Φ : Proof. rewrite wp_unfold /wp_pre. iIntros "[(Hw&HE&Hσ) H]". destruct (to_val e) as [v|] eqn:?; [eauto 10|]. - rewrite fupd_eq. iMod ("H" with "* Hσ [-]") as ">(?&?&%&?)"; first by iFrame. + rewrite fupd_eq. iMod ("H" with "Hσ [-]") as ">(?&?&%&?)"; first by iFrame. eauto 10. Qed. diff --git a/theories/program_logic/ownp.v b/theories/program_logic/ownp.v index 6f08c2b426aa27d104d1019bb60295bb87ed2be0..6d318ea72a2476bcea8a88cdfe74a01d090549d5 100644 --- a/theories/program_logic/ownp.v +++ b/theories/program_logic/ownp.v @@ -96,7 +96,7 @@ Section lifting. iMod (own_update_2 with "Hσ Hσf") as "[Hσ Hσf]". { by apply auth_update, option_local_update, (exclusive_local_update _ (Excl σ2)). } - iFrame "Hσ". iApply ("H" with "* []"); eauto. + iFrame "Hσ". iApply ("H" with "[]"); eauto. Qed. Lemma ownP_lift_pure_step `{Inhabited (state Λ)} E Φ e1 : @@ -171,7 +171,7 @@ Section ectx_lifting. iIntros "H". iApply (ownP_lift_step E); try done. iMod "H" as (σ1) "(%&Hσ1&Hwp)". iModIntro. iExists σ1. iSplit; first by eauto. iFrame. iNext. iIntros (e2 σ2 efs) "% ?". - iApply ("Hwp" with "* []"); by eauto. + iApply ("Hwp" with "[]"); eauto. Qed. Lemma ownP_lift_pure_head_step E Φ e1 : @@ -193,7 +193,7 @@ Section ectx_lifting. ⊢ WP e1 @ E {{ Φ }}. Proof. iIntros (?) "[? H]". iApply ownP_lift_atomic_step; eauto. iFrame. iNext. - iIntros (???) "% ?". iApply ("H" with "* []"); eauto. + iIntros (???) "% ?". iApply ("H" with "[]"); eauto. Qed. Lemma ownP_lift_atomic_det_head_step {E Φ e1} σ1 v2 σ2 efs : diff --git a/theories/program_logic/weakestpre.v b/theories/program_logic/weakestpre.v index 2bdd9fecfb0cc64ac9d603d631662e26de9365ca..70acaa8dc93d95458b595c5c42379d1ab8c73d00 100644 --- a/theories/program_logic/weakestpre.v +++ b/theories/program_logic/weakestpre.v @@ -155,10 +155,10 @@ Proof. { by iDestruct "H" as ">>> $". } iIntros (σ1) "Hσ". iMod "H". iMod ("H" $! σ1 with "Hσ") as "[$ H]". iModIntro. iNext. iIntros (e2 σ2 efs Hstep). - iMod ("H" with "* []") as "(Hphy & H & $)"; first done. + iMod ("H" with "[]") as "(Hphy & H & $)"; first done. rewrite !wp_unfold /wp_pre. destruct (to_val e2) as [v2|] eqn:He2. - iDestruct "H" as ">> $". iFrame. eauto. - - iMod ("H" with "* Hphy") as "[H _]". + - iMod ("H" with "Hphy") as "[H _]". iDestruct "H" as %(? & ? & ? & ?). by edestruct (Hatomic _ _ _ _ Hstep). Qed. diff --git a/theories/proofmode/class_instances.v b/theories/proofmode/class_instances.v index e0f5fdde1f1a8f82a25f39a3a7d0a64e018235a6..0e62d483fa0b29000bfdf7addefd8613e5ccba2a 100644 --- a/theories/proofmode/class_instances.v +++ b/theories/proofmode/class_instances.v @@ -21,6 +21,9 @@ Proof. rewrite /FromAssumption=><-. by rewrite always_always. Qed. Global Instance from_assumption_bupd p P Q : FromAssumption p P Q → FromAssumption p P (|==> Q)%I. Proof. rewrite /FromAssumption=>->. apply bupd_intro. Qed. +Global Instance from_assumption_forall {A} p (Φ : A → uPred M) Q x : + FromAssumption p (Φ x) Q → FromAssumption p (∀ x, Φ x) Q. +Proof. rewrite /FromAssumption=> <-. by rewrite forall_elim. Qed. (* IntoPure *) Global Instance into_pure_pure φ : @IntoPure M ⌜φ⌠φ. @@ -217,6 +220,9 @@ Proof. by apply and_elim_l', impl_wand. Qed. Global Instance into_wand_iff_r P Q : IntoWand (P ↔ Q) Q P. Proof. apply and_elim_r', impl_wand. Qed. +Global Instance into_wand_forall {A} (Φ : A → uPred M) P Q x : + IntoWand (Φ x) P Q → IntoWand (∀ x, Φ x) P Q. +Proof. rewrite /IntoWand=> <-. apply forall_elim. Qed. Global Instance into_wand_always R P Q : IntoWand R P Q → IntoWand (□ R) P Q. Proof. rewrite /IntoWand=> ->. apply always_elim. Qed. diff --git a/theories/proofmode/tactics.v b/theories/proofmode/tactics.v index 3dd911720dfad0094df034741b403a338f92526f..844c3d69a77a08bdefea0fcbc302c7679d29b624 100644 --- a/theories/proofmode/tactics.v +++ b/theories/proofmode/tactics.v @@ -285,7 +285,9 @@ Local Tactic Notation "iSpecializePat" constr(H) constr(pat) := let rec go H1 pats := lazymatch pats with | [] => idtac - | SForall :: ?pats => try (iSpecializeArgs H1 (hcons _ _)); go H1 pats + | SForall :: ?pats => + idtac "the * specialization pattern is deprecated because it is applied implicitly"; + go H1 pats | SName ?H2 :: ?pats => eapply tac_specialize with _ _ H2 _ H1 _ _ _ _; (* (j:=H1) (i:=H2) *) [env_cbv; reflexivity || fail "iSpecialize:" H2 "not found" @@ -424,11 +426,6 @@ Tactic Notation "iPoseProof" open_constr(lem) "as" constr(H) := (** * Apply *) Tactic Notation "iApply" open_constr(lem) := - let lem := (* add a `*` to specialize all top-level foralls *) - lazymatch lem with - | ITrm ?t ?xs ?pat => constr:(ITrm t xs ("*" +:+ pat)) - | _ => constr:(ITrm lem hnil "*") - end in let rec go H := first [eapply tac_apply with _ H _ _ _; [env_cbv; reflexivity @@ -964,27 +961,59 @@ Tactic Notation "iRevertIntros" "(" ident(x1) ident(x2) ident(x3) ident(x4) iRevertIntros (x1 x2 x3 x4 x5 x6 x7 x8) "" with tac. (** * Destruct tactic *) +Class CopyDestruct {M} (P : uPred M). +Hint Mode CopyDestruct + ! : typeclass_instances. + +Instance copy_destruct_forall {M A} (Φ : A → uPred M) : CopyDestruct (∀ x, Φ x). +Instance copy_destruct_impl {M} (P Q : uPred M) : + CopyDestruct Q → CopyDestruct (P → Q). +Instance copy_destruct_wand {M} (P Q : uPred M) : + CopyDestruct Q → CopyDestruct (P -∗ Q). +Instance copy_destruct_always {M} (P : uPred M) : + CopyDestruct P → CopyDestruct (□ P). + Tactic Notation "iDestructCore" open_constr(lem) "as" constr(p) tactic(tac) := + let hyp_name := + lazymatch type of lem with + | string => constr:(Some lem) + | iTrm => + lazymatch lem with + | @iTrm string ?H _ _ => constr:(Some H) | _ => constr:(@None string) + end + | _ => constr:(@None string) + end in let intro_destruct n := let rec go n' := lazymatch n' with | 0 => fail "iDestruct: cannot introduce" n "hypotheses" | 1 => repeat iIntroForall; let H := iFresh in iIntro H; tac H | S ?n' => repeat iIntroForall; let H := iFresh in iIntro H; go n' - end in intros; iStartProof; go n in + end in + intros; iStartProof; go n in lazymatch type of lem with | nat => intro_destruct lem | Z => (* to make it work in Z_scope. We should just be able to bind tactic notation arguments to notation scopes. *) let n := eval compute in (Z.to_nat lem) in intro_destruct n - | string => tac lem - | iTrm => - (* only copy the hypothesis when universal quantifiers are instantiated *) - lazymatch lem with - | @iTrm string ?H _ hnil ?pat => iSpecializeCore lem as p; last tac - | _ => iPoseProofCore lem as p false tac + | _ => + (* Only copy the hypothesis in case there is a [CopyDestruct] instance. + Also, rule out cases in which it does not make sense to copy, namely when + destructing a lemma (instead of a hypothesis) or a spatial hyopthesis + (which cannot be kept). *) + lazymatch hyp_name with + | None => iPoseProofCore lem as p false tac + | Some ?H => iTypeOf H (fun q P => + lazymatch q with + | true => + (* persistent hypothesis, check for a CopyDestruct instance *) + tryif (let dummy := constr:(_ : CopyDestruct P) in idtac) + then (iPoseProofCore lem as p false tac) + else (iSpecializeCore lem as p; last (tac H)) + | false => + (* spatial hypothesis, cannot copy *) + iSpecializeCore lem as p; last (tac H) + end) end - | _ => iPoseProofCore lem as p false tac end. Tactic Notation "iDestruct" open_constr(lem) "as" constr(pat) :=