From f5141d1da510f52b55e7cf6c34a76f727648557b Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Sat, 28 Oct 2017 16:29:32 +0200 Subject: [PATCH] Add `FromForall` instance for `not`. This fixes issue #108. --- theories/proofmode/class_instances.v | 3 +++ theories/tests/proofmode.v | 3 +++ 2 files changed, 6 insertions(+) diff --git a/theories/proofmode/class_instances.v b/theories/proofmode/class_instances.v index 81eef985b..8ba330065 100644 --- a/theories/proofmode/class_instances.v +++ b/theories/proofmode/class_instances.v @@ -772,6 +772,9 @@ Proof. done. Qed. Global Instance from_forall_pure {A} (φ : A → Prop) : @FromForall M A (⌜∀ a : A, φ aâŒ) (λ a, ⌜ φ a âŒ)%I. Proof. by rewrite /FromForall pure_forall. Qed. +Global Instance from_forall_pure_not (φ : Prop) : + @FromForall M φ (⌜¬ φâŒ) (λ a : φ, False)%I. +Proof. by rewrite /FromForall pure_forall. Qed. Global Instance from_forall_impl_pure P Q φ : IntoPureT P φ → FromForall (P → Q) (λ _ : φ, Q)%I. Proof. diff --git a/theories/tests/proofmode.v b/theories/tests/proofmode.v index 1ed83a72f..4aa097847 100644 --- a/theories/tests/proofmode.v +++ b/theories/tests/proofmode.v @@ -233,6 +233,9 @@ Qed. Lemma test_iIntros_rewrite P (x1 x2 x3 x4 : nat) : x1 = x2 → (⌜ x2 = x3 ⌠∗ ⌜ x3 ≡ x4 ⌠∗ P) -∗ ⌜ x1 = x4 ⌠∗ P. Proof. iIntros (?) "(-> & -> & $)"; auto. Qed. + +Lemma test_iItros_pure : (⌜ ¬False ⌠: uPred M)%I. +Proof. by iIntros (?). Qed. End tests. Section more_tests. -- GitLab