feat(library/standard): add or_comm, and_comm, ... theorems, cleanup notation

Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
This commit is contained in:
Leonardo de Moura 2014-07-19 09:29:04 +01:00
parent 66ba3c8a0b
commit e817260c6d
4 changed files with 127 additions and 102 deletions

View file

@ -46,12 +46,12 @@ theorem not_not_elim {a : Bool} (H : ¬ ¬ a) : a
theorem boolext {a b : Bool} (Hab : a → b) (Hba : b → a) : a = b theorem boolext {a b : Bool} (Hab : a → b) (Hba : b → a) : a = b
:= or_elim (boolcomplete a) := or_elim (boolcomplete a)
(λ Hat : a = true, or_elim (boolcomplete b) (assume Hat, or_elim (boolcomplete b)
(λ Hbt : b = true, trans Hat (symm Hbt)) (assume Hbt, trans Hat (symm Hbt))
(λ Hbf : b = false, false_elim (a = b) (subst Hbf (Hab (eqt_elim Hat))))) (assume Hbf, false_elim (a = b) (subst Hbf (Hab (eqt_elim Hat)))))
(λ Haf : a = false, or_elim (boolcomplete b) (assume Haf, or_elim (boolcomplete b)
(λ Hbt : b = true, false_elim (a = b) (subst Haf (Hba (eqt_elim Hbt)))) (assume Hbt, false_elim (a = b) (subst Haf (Hba (eqt_elim Hbt))))
(λ Hbf : b = false, trans Haf (symm Hbf))) (assume Hbf, trans Haf (symm Hbf)))
theorem iff_to_eq {a b : Bool} (H : a ↔ b) : a = b theorem iff_to_eq {a b : Bool} (H : a ↔ b) : a = b
:= iff_elim (assume H1 H2, boolext H1 H2) H := iff_elim (assume H1 H2, boolext H1 H2) H
@ -70,7 +70,7 @@ theorem eqf_intro {a : Bool} (H : ¬ a) : a = false
(assume H2 : false, false_elim a H2) (assume H2 : false, false_elim a H2)
theorem by_contradiction {a : Bool} (H : ¬a → false) : a theorem by_contradiction {a : Bool} (H : ¬a → false) : a
:= or_elim (em a) (λ H1 : a, H1) (λ H1 : ¬ a, false_elim a (H H1)) := or_elim (em a) (assume H1 : a, H1) (assume H1 : ¬a, false_elim a (H H1))
theorem a_neq_a {A : Type} (a : A) : (a ≠ a) = false theorem a_neq_a {A : Type} (a : A) : (a ≠ a) = false
:= boolext (assume H, a_neq_a_elim H) := boolext (assume H, a_neq_a_elim H)
@ -108,12 +108,10 @@ theorem not_and (a b : Bool) : (¬ (a ∧ b)) = (¬ a ¬ b)
theorem imp_or (a b : Bool) : (a → b) = (¬ a b) theorem imp_or (a b : Bool) : (a → b) = (¬ a b)
:= boolext := boolext
(assume H : a → b, (assume H : a → b, (or_elim (em a)
(or_elim (em a) (assume Ha : a, or_intro_right (¬ a) (H Ha))
(λ Ha : a, or_intro_right (¬ a) (H Ha)) (assume Hna : ¬ a, or_intro_left b Hna)))
(λ Hna : ¬ a, or_intro_left b Hna))) (assume (H : ¬ a b) (Ha : a),
(assume H : ¬ a b,
assume Ha : a,
resolve_right H ((symm (not_not_eq a)) ◂ Ha)) resolve_right H ((symm (not_not_eq a)) ◂ Ha))
theorem not_implies (a b : Bool) : (¬ (a → b)) = (a ∧ ¬b) theorem not_implies (a b : Bool) : (¬ (a → b)) = (a ∧ ¬b)
@ -135,14 +133,10 @@ theorem false_eq_true : (false = true) = false
:= subst not_false (a_eq_not_a false) := subst not_false (a_eq_not_a false)
theorem a_eq_true (a : Bool) : (a = true) = a theorem a_eq_true (a : Bool) : (a = true) = a
:= boolext := boolext (assume H, eqt_elim H) (assume H, eqt_intro H)
(assume H, eqt_elim H)
(assume H, eqt_intro H)
theorem a_eq_false (a : Bool) : (a = false) = (¬ a) theorem a_eq_false (a : Bool) : (a = false) = ¬a
:= boolext := boolext (assume H, eqf_elim H) (assume H, eqf_intro H)
(assume H, eqf_elim H)
(assume H, eqf_intro H)
theorem not_exists_forall {A : Type} {P : A → Bool} (H : ¬∃x, P x) : ∀x, ¬P x theorem not_exists_forall {A : Type} {P : A → Bool} (H : ¬∃x, P x) : ∀x, ¬P x
:= take x, or_elim (em (P x)) := take x, or_elim (em (P x))
@ -156,6 +150,6 @@ theorem not_forall_exists {A : Type} {P : A → Bool} (H : ¬ ∀ x, P x) : ∃
absurd H3 H) absurd H3 H)
theorem peirce (a b : Bool) : ((a → b) → a) → a theorem peirce (a b : Bool) : ((a → b) → a) → a
:= assume H, by_contradiction (λ Hna : ¬ a, := assume H, by_contradiction (assume Hna : ¬a,
have Hnna : ¬¬a, from not_implies_left (mt H Hna), have Hnna : ¬¬a, from not_implies_left (mt H Hna),
absurd (not_not_elim Hnna) Hna) absurd (not_not_elim Hnna) Hna)

View file

@ -1,6 +1,6 @@
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved. -- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura -- Authors: Leonardo de Moura, Jeremy Avigad
definition Bool [inline] := Type.{0} definition Bool [inline] := Type.{0}
inductive false : Bool := inductive false : Bool :=
@ -63,6 +63,9 @@ theorem and_elim_left {a b : Bool} (H : a ∧ b) : a
theorem and_elim_right {a b : Bool} (H : a ∧ b) : b theorem and_elim_right {a b : Bool} (H : a ∧ b) : b
:= and_rec (λa b, b) H := and_rec (λa b, b) H
theorem and_swap {a b : Bool} (H : a ∧ b) : b ∧ a
:= and_intro (and_elim_right H) (and_elim_left H)
inductive or (a b : Bool) : Bool := inductive or (a b : Bool) : Bool :=
| or_intro_left : a → or a b | or_intro_left : a → or a b
| or_intro_right : b → or a b | or_intro_right : b → or a b
@ -79,7 +82,7 @@ theorem resolve_right {a b : Bool} (H1 : a b) (H2 : ¬ a) : b
theorem resolve_left {a b : Bool} (H1 : a b) (H2 : ¬b) : a theorem resolve_left {a b : Bool} (H1 : a b) (H2 : ¬b) : a
:= or_elim H1 (assume Ha, Ha) (assume Hb, absurd_elim a Hb H2) := or_elim H1 (assume Ha, Ha) (assume Hb, absurd_elim a Hb H2)
theorem or_flip {a b : Bool} (H : a b) : b a theorem or_swap {a b : Bool} (H : a b) : b a
:= or_elim H (assume Ha, or_intro_right b Ha) (assume Hb, or_intro_left a Hb) := or_elim H (assume Ha, or_intro_right b Ha) (assume Hb, or_intro_left a Hb)
inductive eq {A : Type} (a : A) : A → Bool := inductive eq {A : Type} (a : A) : A → Bool :=
@ -171,7 +174,7 @@ calc_trans eq_ne_trans
calc_trans ne_eq_trans calc_trans ne_eq_trans
definition iff (a b : Bool) := (a → b) ∧ (b → a) definition iff (a b : Bool) := (a → b) ∧ (b → a)
infix `↔`:50 := iff infix `↔`:25 := iff
theorem iff_intro {a b : Bool} (H1 : a → b) (H2 : b → a) : a ↔ b theorem iff_intro {a b : Bool} (H1 : a → b) (H2 : b → a) : a ↔ b
:= and_intro H1 H2 := and_intro H1 H2
@ -194,22 +197,50 @@ theorem iff_mp_right {a b : Bool} (H1 : a ↔ b) (H2 : b) : a
theorem eq_to_iff {a b : Bool} (H : a = b) : a ↔ b theorem eq_to_iff {a b : Bool} (H : a = b) : a ↔ b
:= iff_intro (λ Ha, subst H Ha) (λ Hb, subst (symm H) Hb) := iff_intro (λ Ha, subst H Ha) (λ Hb, subst (symm H) Hb)
theorem and_comm (a b : Bool) : a ∧ b ↔ b ∧ a
:= iff_intro (λH, and_swap H) (λH, and_swap H)
theorem and_assoc (a b c : Bool) : (a ∧ b) ∧ c ↔ a ∧ (b ∧ c)
:= iff_intro
(assume H, and_intro
(and_elim_left (and_elim_left H))
(and_intro (and_elim_right (and_elim_left H)) (and_elim_right H)))
(assume H, and_intro
(and_intro (and_elim_left H) (and_elim_left (and_elim_right H)))
(and_elim_right (and_elim_right H)))
theorem or_comm (a b : Bool) : a b ↔ b a
:= iff_intro (λH, or_swap H) (λH, or_swap H)
theorem or_assoc (a b c : Bool) : (a b) c ↔ a (b c)
:= iff_intro
(assume H, or_elim H
(assume H1, or_elim H1
(assume Ha, or_intro_left _ Ha)
(assume Hb, or_intro_right a (or_intro_left c Hb)))
(assume Hc, or_intro_right a (or_intro_right b Hc)))
(assume H, or_elim H
(assume Ha, (or_intro_left c (or_intro_left b Ha)))
(assume H1, or_elim H1
(assume Hb, or_intro_left c (or_intro_right a Hb))
(assume Hc, or_intro_right _ Hc)))
inductive Exists {A : Type} (P : A → Bool) : Bool := inductive Exists {A : Type} (P : A → Bool) : Bool :=
| exists_intro : ∀ (a : A), P a → Exists P | exists_intro : ∀ (a : A), P a → Exists P
notation `∃` binders `,` r:(scoped P, Exists P) := r notation `∃` binders `,` r:(scoped P, Exists P) := r
theorem exists_elim {A : Type} {P : A → Bool} {B : Bool} (H1 : ∃ x : A, P x) (H2 : ∀ (a : A) (H : P a), B) : B theorem exists_elim {A : Type} {p : A → Bool} {B : Bool} (H1 : ∃x, p x) (H2 : ∀ (a : A) (H : p a), B) : B
:= Exists_rec H2 H1 := Exists_rec H2 H1
theorem exists_not_forall {A : Type} {P : A → Bool} (H : ∃ x, P x) : ¬ ∀ x, ¬ P x theorem exists_not_forall {A : Type} {p : A → Bool} (H : ∃x, p x) : ¬∀x, ¬p x
:= assume H1 : ∀ x, ¬ P x, := assume H1 : ∀x, ¬p x,
obtain (w : A) (Hw : P w), from H, obtain (w : A) (Hw : p w), from H,
absurd Hw (H1 w) absurd Hw (H1 w)
theorem forall_not_exists {A : Type} {P : A → Bool} (H2 : ∀ x, P x) : ¬ ∃ x, ¬ P x theorem forall_not_exists {A : Type} {p : A → Bool} (H2 : ∀x, p x) : ¬∃x, ¬p x
:= assume H1 : ∃ x, ¬ P x, := assume H1 : ∃x, ¬p x,
obtain (w : A) (Hw : ¬ P w), from H1, obtain (w : A) (Hw : ¬p w), from H1,
absurd (H2 w) Hw absurd (H2 w) Hw
definition exists_unique {A : Type} (p : A → Bool) := ∃x, p x ∧ ∀y, y ≠ x → ¬ p y definition exists_unique {A : Type} (p : A → Bool) := ∃x, p x ∧ ∀y, y ≠ x → ¬ p y
@ -233,7 +264,7 @@ theorem inhabited_Bool [instance] : inhabited Bool
:= inhabited_intro true := inhabited_intro true
theorem inhabited_fun [instance] (A : Type) {B : Type} (H : inhabited B) : inhabited (A → B) theorem inhabited_fun [instance] (A : Type) {B : Type} (H : inhabited B) : inhabited (A → B)
:= inhabited_elim H (take (b : B), inhabited_intro (λ a : A, b)) := inhabited_elim H (take b, inhabited_intro (λa, b))
theorem inhabited_exists {A : Type} {P : A → Bool} (H : ∃ x, P x) : inhabited A theorem inhabited_exists {A : Type} {p : A → Bool} (H : ∃x, p x) : inhabited A
:= obtain w Hw, from H, inhabited_intro w := obtain w Hw, from H, inhabited_intro w

View file

@ -16,7 +16,7 @@ theorem wf_induction {A : Type} {R : A → A → Bool} {P : A → Bool} (Hwf : w
obtain (w : A) (Hw : ¬P w), from not_forall_exists N, obtain (w : A) (Hw : ¬P w), from not_forall_exists N,
-- The main "trick" is to define Q x as ¬ P x. -- The main "trick" is to define Q x as ¬ P x.
-- Since R is well-founded, there must be a R-minimal element r s.t. Q r (which is ¬ P r) -- Since R is well-founded, there must be a R-minimal element r s.t. Q r (which is ¬ P r)
let Q [inline] := λ x, ¬ P x in let Q [inline] x := ¬P x in
have Qw : ∃w, Q w, from exists_intro w Hw, have Qw : ∃w, Q w, from exists_intro w Hw,
have Qwf : ∃min, Q min ∧ ∀b, R b min → ¬Q b, from Hwf Q Qw, have Qwf : ∃min, Q min ∧ ∀b, R b min → ¬Q b, from Hwf Q Qw,
obtain (r : A) (Hr : Q r ∧ ∀b, R b r → ¬Q b), from Qwf, obtain (r : A) (Hr : Q r ∧ ∀b, R b r → ¬Q b), from Qwf,
@ -27,4 +27,4 @@ theorem wf_induction {A : Type} {R : A → A → Bool} {P : A → Bool} (Hwf : w
not_not_elim (and_elim_right Hr b H), not_not_elim (and_elim_right Hr b H),
have s2 : P r, from iH r s1, have s2 : P r, from iH r s1,
have s3 : ¬P r, from and_elim_left Hr, have s3 : ¬P r, from and_elim_left Hr,
show false, from absurd s2 s3) absurd s2 s3)