2014-12-15 21:13:04 +00:00
|
|
|
|
/-
|
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
Authors: Jeremy Avigad, Leonardo de Moura
|
2014-08-22 23:36:47 +00:00
|
|
|
|
|
2014-12-15 21:13:04 +00:00
|
|
|
|
Useful logical identities. Since we are not using propositional extensionality, some of the
|
|
|
|
|
calculations use the type class support provided by logic.instances.
|
|
|
|
|
-/
|
2015-11-21 00:38:10 +00:00
|
|
|
|
import logic.connectives logic.quantifiers logic.cast
|
|
|
|
|
open decidable
|
2014-08-22 23:36:47 +00:00
|
|
|
|
|
2014-12-15 21:13:04 +00:00
|
|
|
|
theorem or.right_comm (a b c : Prop) : (a ∨ b) ∨ c ↔ (a ∨ c) ∨ b :=
|
2014-08-22 23:36:47 +00:00
|
|
|
|
calc
|
2014-09-05 04:25:21 +00:00
|
|
|
|
(a ∨ b) ∨ c ↔ a ∨ (b ∨ c) : or.assoc
|
|
|
|
|
... ↔ a ∨ (c ∨ b) : {or.comm}
|
|
|
|
|
... ↔ (a ∨ c) ∨ b : iff.symm or.assoc
|
2014-08-22 23:36:47 +00:00
|
|
|
|
|
2014-12-15 21:13:04 +00:00
|
|
|
|
theorem and.right_comm (a b c : Prop) : (a ∧ b) ∧ c ↔ (a ∧ c) ∧ b :=
|
2014-08-22 23:36:47 +00:00
|
|
|
|
calc
|
2014-09-05 04:25:21 +00:00
|
|
|
|
(a ∧ b) ∧ c ↔ a ∧ (b ∧ c) : and.assoc
|
|
|
|
|
... ↔ a ∧ (c ∧ b) : {and.comm}
|
|
|
|
|
... ↔ (a ∧ c) ∧ b : iff.symm and.assoc
|
2014-08-22 23:36:47 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem or_not_self_iff (a : Prop) [D : decidable a] : a ∨ ¬ a ↔ true :=
|
2015-09-25 03:38:52 +00:00
|
|
|
|
iff.intro (assume H, trivial) (assume H, em a)
|
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_or_self_iff (a : Prop) [D : decidable a] : ¬ a ∨ a ↔ true :=
|
2015-11-21 00:38:10 +00:00
|
|
|
|
iff.intro (λ H, trivial) (λ H, or.swap (em a))
|
2015-09-25 03:38:52 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem and_not_self_iff (a : Prop) : a ∧ ¬ a ↔ false :=
|
2015-09-25 03:38:52 +00:00
|
|
|
|
iff.intro (assume H, (and.right H) (and.left H)) (assume H, false.elim H)
|
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_and_self_iff (a : Prop) : ¬ a ∧ a ↔ false :=
|
2015-11-21 00:38:10 +00:00
|
|
|
|
iff.intro (λ H, and.elim H (by contradiction)) (λ H, false.elim H)
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_not_iff (a : Prop) [D : decidable a] : ¬¬a ↔ a :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
iff.intro by_contradiction not_not_intro
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2015-07-24 15:56:18 +00:00
|
|
|
|
theorem not_not_elim {a : Prop} [D : decidable a] : ¬¬a → a :=
|
|
|
|
|
by_contradiction
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_or_iff_not_and_not (a b : Prop) : ¬(a ∨ b) ↔ ¬a ∧ ¬b :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
or.imp_distrib
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_or_not_of_not_and {a b : Prop} [Da : decidable a] (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b :=
|
|
|
|
|
by_cases (λHa, or.inr (not.mto (and.intro Ha) H)) or.inl
|
|
|
|
|
|
|
|
|
|
theorem not_or_not_of_not_and' {a b : Prop} [Db : decidable b] (H : ¬ (a ∧ b)) : ¬ a ∨ ¬ b :=
|
|
|
|
|
by_cases (λHb, or.inl (λHa, H (and.intro Ha Hb))) or.inr
|
|
|
|
|
|
|
|
|
|
theorem not_and_iff_not_or_not (a b : Prop) [Da : decidable a] :
|
2014-12-15 21:13:04 +00:00
|
|
|
|
¬(a ∧ b) ↔ ¬a ∨ ¬b :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
2016-02-16 16:54:26 +00:00
|
|
|
|
not_or_not_of_not_and
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(or.rec (not.mto and.left) (not.mto and.right))
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem or_iff_not_and_not (a b : Prop) [Da : decidable a] [Db : decidable b] :
|
2015-08-27 18:26:21 +00:00
|
|
|
|
a ∨ b ↔ ¬ (¬a ∧ ¬b) :=
|
|
|
|
|
by rewrite [-not_or_iff_not_and_not, not_not_iff]
|
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem and_iff_not_or_not (a b : Prop) [Da : decidable a] [Db : decidable b] :
|
2015-08-27 18:26:21 +00:00
|
|
|
|
a ∧ b ↔ ¬ (¬ a ∨ ¬ b) :=
|
|
|
|
|
by rewrite [-not_and_iff_not_or_not, not_not_iff]
|
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem imp_iff_not_or (a b : Prop) [Da : decidable a] : (a → b) ↔ ¬a ∨ b :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
2015-07-24 15:56:18 +00:00
|
|
|
|
(by_cases (λHa H, or.inr (H Ha)) (λHa H, or.inl Ha))
|
|
|
|
|
(or.rec not.elim imp.intro)
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem not_implies_iff_and_not (a b : Prop) [Da : decidable a] :
|
2014-12-15 21:13:04 +00:00
|
|
|
|
¬(a → b) ↔ a ∧ ¬b :=
|
|
|
|
|
calc
|
2016-02-16 16:54:26 +00:00
|
|
|
|
¬(a → b) ↔ ¬(¬a ∨ b) : {imp_iff_not_or a b}
|
2014-12-15 21:13:04 +00:00
|
|
|
|
... ↔ ¬¬a ∧ ¬b : not_or_iff_not_and_not
|
2016-02-16 16:54:26 +00:00
|
|
|
|
... ↔ a ∧ ¬b : {not_not_iff a}
|
|
|
|
|
|
|
|
|
|
theorem and_not_of_not_implies {a b : Prop} [Da : decidable a] (H : ¬ (a → b)) : a ∧ ¬ b :=
|
|
|
|
|
iff.mp !not_implies_iff_and_not H
|
|
|
|
|
|
|
|
|
|
theorem not_implies_of_and_not {a b : Prop} [Da : decidable a] (H : a ∧ ¬ b) : ¬ (a → b) :=
|
|
|
|
|
iff.mpr !not_implies_iff_and_not H
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2016-02-16 16:54:26 +00:00
|
|
|
|
theorem peirce (a b : Prop) [D : decidable a] : ((a → b) → a) → a :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
by_cases imp.intro (imp.syl imp.mp not.elim)
|
2014-08-30 03:45:57 +00:00
|
|
|
|
|
2015-06-04 03:41:52 +00:00
|
|
|
|
theorem forall_not_of_not_exists {A : Type} {p : A → Prop} [D : ∀x, decidable (p x)]
|
|
|
|
|
(H : ¬∃x, p x) : ∀x, ¬p x :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
take x, by_cases
|
|
|
|
|
(assume Hp : p x, absurd (exists.intro x Hp) H)
|
|
|
|
|
imp.id
|
2015-06-04 03:41:52 +00:00
|
|
|
|
|
|
|
|
|
theorem forall_of_not_exists_not {A : Type} {p : A → Prop} [D : decidable_pred p] :
|
|
|
|
|
¬(∃ x, ¬p x) → ∀ x, p x :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
imp.syl (forall_imp_forall (λa, not_not_elim)) forall_not_of_not_exists
|
2015-06-04 03:41:52 +00:00
|
|
|
|
|
|
|
|
|
theorem exists_not_of_not_forall {A : Type} {p : A → Prop} [D : ∀x, decidable (p x)]
|
|
|
|
|
[D' : decidable (∃x, ¬p x)] (H : ¬∀x, p x) :
|
|
|
|
|
∃x, ¬p x :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
by_contradiction (λH1, absurd (λx, not_not_elim (forall_not_of_not_exists H1 x)) H)
|
2015-06-04 03:41:52 +00:00
|
|
|
|
|
|
|
|
|
theorem exists_of_not_forall_not {A : Type} {p : A → Prop} [D : ∀x, decidable (p x)]
|
2015-07-24 15:56:18 +00:00
|
|
|
|
[D' : decidable (∃x, p x)] (H : ¬∀x, ¬ p x) :
|
2015-06-04 03:41:52 +00:00
|
|
|
|
∃x, p x :=
|
2015-07-24 15:56:18 +00:00
|
|
|
|
by_contradiction (imp.syl H forall_not_of_not_exists)
|