feat(library/logic/axioms/prop_complete): add by_cases, by_contradiction

This commit is contained in:
Jeremy Avigad 2015-02-01 11:16:30 -05:00 committed by Leonardo de Moura
parent 003a2c1e2c
commit 5ef510f290
2 changed files with 12 additions and 7 deletions

View file

@ -5,9 +5,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Module: logic.axioms.classical Module: logic.axioms.classical
Author: Leonardo de Moura Author: Leonardo de Moura
-/ -/
import logic.connectives logic.quantifiers logic.cast algebra.relation import logic.connectives logic.quantifiers logic.cast algebra.relation
open eq.ops open eq.ops
axiom prop_complete (a : Prop) : a = true a = false axiom prop_complete (a : Prop) : a = true a = false
@ -28,6 +26,16 @@ or.elim (prop_complete a)
(assume Ht : a = true, or.inl (of_eq_true Ht)) (assume Ht : a = true, or.inl (of_eq_true Ht))
(assume Hf : a = false, or.inr (not_of_eq_false Hf)) (assume Hf : a = false, or.inr (not_of_eq_false Hf))
-- this supercedes by_cases in decidable
definition by_cases {p q : Prop} (Hpq : p → q) (Hnpq : ¬p → q) : q :=
or.elim (em p) (assume Hp, Hpq Hp) (assume Hnp, Hnpq Hnp)
-- this supercedes by_contradiction in decidable
theorem by_contradiction {p : Prop} (H : ¬p → false) : p :=
by_cases
(assume H1 : p, H1)
(assume H1 : ¬p, false.rec _ (H H1))
theorem eq_false_or_eq_true (a : Prop) : a = false a = true := theorem eq_false_or_eq_true (a : Prop) : a = false a = true :=
cases (λ x, x = false x = true) cases (λ x, x = false x = true)
(or.inr rfl) (or.inr rfl)

View file

@ -8,9 +8,7 @@ Authors: Jeremy Avigad, Leonardo de Moura
Useful logical identities. Since we are not using propositional extensionality, some of the Useful logical identities. Since we are not using propositional extensionality, some of the
calculations use the type class support provided by logic.instances. calculations use the type class support provided by logic.instances.
-/ -/
import logic.connectives logic.instances logic.quantifiers logic.cast import logic.connectives logic.instances logic.quantifiers logic.cast
open relation decidable relation.iff_ops open relation decidable relation.iff_ops
theorem or.right_comm (a b c : Prop) : (a b) c ↔ (a c) b := theorem or.right_comm (a b c : Prop) : (a b) c ↔ (a c) b :=
@ -19,7 +17,7 @@ calc
... ↔ a (c b) : {or.comm} ... ↔ a (c b) : {or.comm}
... ↔ (a c) b : iff.symm or.assoc ... ↔ (a c) b : iff.symm or.assoc
theorem or.left_comm (a b c : Prop) : a (b c)↔ b (a c) := theorem or.left_comm (a b c : Prop) : a (b c) ↔ b (a c) :=
calc calc
a (b c) ↔ (a b) c : iff.symm or.assoc a (b c) ↔ (a b) c : iff.symm or.assoc
... ↔ (b a) c : {or.comm} ... ↔ (b a) c : {or.comm}
@ -31,13 +29,12 @@ calc
... ↔ a ∧ (c ∧ b) : {and.comm} ... ↔ a ∧ (c ∧ b) : {and.comm}
... ↔ (a ∧ c) ∧ b : iff.symm and.assoc ... ↔ (a ∧ c) ∧ b : iff.symm and.assoc
theorem and.left_comm (a b c : Prop) : a ∧ (b ∧ c)↔ b ∧ (a ∧ c) := theorem and.left_comm (a b c : Prop) : a ∧ (b ∧ c) ↔ b ∧ (a ∧ c) :=
calc calc
a ∧ (b ∧ c) ↔ (a ∧ b) ∧ c : iff.symm and.assoc a ∧ (b ∧ c) ↔ (a ∧ b) ∧ c : iff.symm and.assoc
... ↔ (b ∧ a) ∧ c : {and.comm} ... ↔ (b ∧ a) ∧ c : {and.comm}
... ↔ b ∧ (a ∧ c) : and.assoc ... ↔ b ∧ (a ∧ c) : and.assoc
theorem not_not_iff {a : Prop} [D : decidable a] : (¬¬a) ↔ a := theorem not_not_iff {a : Prop} [D : decidable a] : (¬¬a) ↔ a :=
iff.intro iff.intro
(assume H : ¬¬a, (assume H : ¬¬a,