From 2285a577625e4a14834e59e0abc65ce788d05a8d Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Thu, 4 Aug 2016 22:34:47 +0200 Subject: [PATCH] Make iApply more powerful and uniform. MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit It not behaves more consistently with iExact and thus also works in the case H : P -★ □^n Q |- Q. --- heap_lang/lib/barrier/proof.v | 2 +- program_logic/counter_examples.v | 55 ++++++++++++-------------------- proofmode/class_instances.v | 10 +++--- proofmode/pviewshifts.v | 7 ++-- 4 files changed, 31 insertions(+), 43 deletions(-) diff --git a/heap_lang/lib/barrier/proof.v b/heap_lang/lib/barrier/proof.v index 879ac67ca..ba12947ad 100644 --- a/heap_lang/lib/barrier/proof.v +++ b/heap_lang/lib/barrier/proof.v @@ -99,7 +99,7 @@ Lemma newbarrier_spec (P : iProp) (Φ : val → iProp) : Proof. iIntros (HN) "[#? HΦ]". rewrite /newbarrier. wp_seq. wp_alloc l as "Hl". - iApply "HΦ". + iApply ("HΦ" with "|==>[-]"). iPvs (saved_prop_alloc (F:=idCF) _ P) as (γ) "#?". iPvs (sts_alloc (barrier_inv l P) _ N (State Low {[ γ ]}) with "[-]") as (γ') "[#? Hγ']"; eauto. diff --git a/program_logic/counter_examples.v b/program_logic/counter_examples.v index 1aabfbea3..14710eba5 100644 --- a/program_logic/counter_examples.v +++ b/program_logic/counter_examples.v @@ -7,66 +7,51 @@ From iris.proofmode Require Import tactics. Section savedprop. Context (M : ucmraT). Notation iProp := (uPred M). - Notation "¬ P" := (□(P → False))%I : uPred_scope. + Notation "¬ P" := (□ (P → False))%I : uPred_scope. + Implicit Types P : iProp. (* Saved Propositions and view shifts. *) Context (sprop : Type) (saved : sprop → iProp → iProp) (pvs : iProp → iProp). - Hypothesis pvs_mono : forall P Q, (P ⊢ Q) → pvs P ⊢ pvs Q. - Hypothesis sprop_persistent : forall i P, PersistentP (saved i P). + Hypothesis pvs_mono : ∀ P Q, (P ⊢ Q) → pvs P ⊢ pvs Q. + Hypothesis sprop_persistent : ∀ i P, PersistentP (saved i P). Hypothesis sprop_alloc_dep : - forall (P : sprop → iProp), True ⊢ pvs (∃ i, saved i (P i)). - Hypothesis sprop_agree : - forall i P Q, saved i P ∧ saved i Q ⊢ P ↔ Q. + ∀ (P : sprop → iProp), True ⊢ pvs (∃ i, saved i (P i)). + Hypothesis sprop_agree : ∀ i P Q, saved i P ∧ saved i Q ⊢ P ↔ Q. (* Self-contradicting assertions are inconsistent *) - Lemma no_self_contradiction (P : iProp) `{!PersistentP P} : - □(P ↔ ¬ P) ⊢ False. + Lemma no_self_contradiction P `{!PersistentP P} : □ (P ↔ ¬ P) ⊢ False. Proof. (* FIXME: Cannot destruct the <-> as two implications. iApply with <-> also does not work. *) - rewrite /uPred_iff. iIntros "#[H1 H2]". (* FIXME: Cannot iApply "H1". *) + rewrite /uPred_iff. iIntros "#[H1 H2]". iAssert P as "#HP". - { iApply "H2". iIntros "!#HP". by iApply ("H1" with "HP"). } - by iApply ("H1" with "HP HP"). + { iApply "H2". iIntros "! #HP". by iApply ("H1" with "[#]"). } + by iApply ("H1" with "[#]"). Qed. (* "Assertion with name [i]" is equivalent to any assertion P s.t. [saved i P] *) - Definition A (i : sprop) : iProp := ∃ P, saved i P ★ □P. - Lemma saved_is_A i P `{!PersistentP P} : saved i P ⊢ □(A i ↔ P). + Definition A (i : sprop) : iProp := ∃ P, saved i P ★ □ P. + Lemma saved_is_A i P `{!PersistentP P} : saved i P ⊢ □ (A i ↔ P). Proof. rewrite /uPred_iff. iIntros "#HS !". iSplit. - - iIntros "H". iDestruct "H" as (Q) "[#HSQ HQ]". - iPoseProof (sprop_agree i P Q with "[]") as "Heq"; first by eauto. - rewrite /uPred_iff. by iApply "Heq". + - iDestruct 1 as (Q) "[#HSQ HQ]". + iApply (sprop_agree i P Q with "[]"); eauto. - iIntros "#HP". iExists P. by iSplit. Qed. (* Define [Q i] to be "negated assertion with name [i]". Show that this implies that assertion with name [i] is equivalent to its own negation. *) Definition Q i := saved i (¬ A i). - Lemma Q_self_contradiction i : - Q i ⊢ □(A i ↔ ¬ A i). - Proof. - iIntros "#HQ". iApply (@saved_is_A i (¬ A i)%I _). (* FIXME: If we already introduced the box, this iApply does not work. *) - done. - Qed. + Lemma Q_self_contradiction i : Q i ⊢ □ (A i ↔ ¬ A i). + Proof. iIntros "#HQ !". by iApply (saved_is_A i (¬A i)). Qed. (* We can obtain such a [Q i]. *) - Lemma make_Q : - True ⊢ pvs (∃ i, Q i). - Proof. - apply sprop_alloc_dep. - Qed. + Lemma make_Q : True ⊢ pvs (∃ i, Q i). + Proof. apply sprop_alloc_dep. Qed. (* Put together all the pieces to derive a contradiction. *) (* TODO: Have a lemma in upred.v that says that we cannot view shift to False. *) Lemma contradiction : True ⊢ pvs False. Proof. - rewrite make_Q. apply pvs_mono. - iIntros "HQ". iDestruct "HQ" as (i) "HQ". - iApply (@no_self_contradiction (A i) _). - by iApply Q_self_contradiction. + rewrite make_Q. apply pvs_mono. iDestruct 1 as (i) "HQ". + iApply (no_self_contradiction (A i)). by iApply Q_self_contradiction. Qed. End savedprop. - - - - diff --git a/proofmode/class_instances.v b/proofmode/class_instances.v index dfa04ffe4..fede740a3 100644 --- a/proofmode/class_instances.v +++ b/proofmode/class_instances.v @@ -87,10 +87,12 @@ Global Instance from_later_sep P1 P2 Q1 Q2 : Proof. intros ??; red. by rewrite later_sep; apply sep_mono. Qed. (* IntoWand *) -Global Instance into_wand_wand P Q : IntoWand (P -★ Q) P Q. -Proof. done. Qed. -Global Instance into_wand_impl P Q : IntoWand (P → Q) P Q. -Proof. apply impl_wand. Qed. +Global Instance into_wand_wand P Q Q' : + FromAssumption false Q Q' → IntoWand (P -★ Q) P Q'. +Proof. by rewrite /FromAssumption /IntoWand /= => ->. Qed. +Global Instance into_wand_impl P Q Q' : + FromAssumption false Q Q' → IntoWand (P → Q) P Q'. +Proof. rewrite /FromAssumption /IntoWand /= => ->. by rewrite impl_wand. Qed. Global Instance into_wand_iff_l P Q : IntoWand (P ↔ Q) P Q. Proof. by apply and_elim_l', impl_wand. Qed. Global Instance into_wand_iff_r P Q : IntoWand (P ↔ Q) Q P. diff --git a/proofmode/pviewshifts.v b/proofmode/pviewshifts.v index 676710d71..2de2ce5f9 100644 --- a/proofmode/pviewshifts.v +++ b/proofmode/pviewshifts.v @@ -12,6 +12,10 @@ Proof. intros ??. by rewrite -pvs_intro (from_pure P). Qed. Global Instance from_assumption_pvs E p P Q : FromAssumption p P Q → FromAssumption p P (|={E}=> Q)%I. Proof. rewrite /FromAssumption=>->. apply pvs_intro. Qed. +Global Instance into_wand_pvs E1 E2 R P Q : + IntoWand R P Q → IntoWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100. +Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed. + Global Instance from_sep_pvs E P Q1 Q2 : FromSep P Q1 Q2 → FromSep (|={E}=> P) (|={E}=> Q1) (|={E}=> Q2). Proof. rewrite /FromSep=><-. apply pvs_sep. Qed. @@ -26,9 +30,6 @@ Qed. Global Instance frame_pvs E1 E2 R P Q : Frame R P Q → Frame R (|={E1,E2}=> P) (|={E1,E2}=> Q). Proof. rewrite /Frame=><-. by rewrite pvs_frame_l. Qed. -Global Instance into_wand_pvs E1 E2 R P Q : - IntoWand R P Q → IntoWand R (|={E1,E2}=> P) (|={E1,E2}=> Q) | 100. -Proof. rewrite /IntoWand=>->. apply wand_intro_l. by rewrite pvs_wand_r. Qed. Global Instance timeless_elim_pvs E1 E2 Q : TimelessElim (|={E1,E2}=> Q). Proof. -- GitLab