feat(library/logic/{connectives,identities},library/algebra/function): cleanup and some additions from Haitao Zhang

This commit is contained in:
Jeremy Avigad 2015-06-04 13:41:52 +10:00 committed by Leonardo de Moura
parent 134740182d
commit 03952ae12c
3 changed files with 43 additions and 14 deletions

View file

@ -37,7 +37,7 @@ definition dcompose [reducible] [unfold-f] {B : A → Type} {C : Π {x : A}, B x
(f : Π {x : A} (y : B x), C y) (g : Πx, B x) : Πx, C (g x) :=
λx, f (g x)
definition flip [reducible] [unfold-f] {C : A → B → Type} (f : Πx y, C x y) : Πy x, C x y :=
definition swap [reducible] [unfold-f] {C : A → B → Type} (f : Πx y, C x y) : Πy x, C x y :=
λy x, f x y
definition app [reducible] {B : A → Type} (f : Πx, B x) (x : A) : B x :=

View file

@ -1,7 +1,7 @@
/-
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
Authors: Jeremy Avigad, Leonardo de Moura, Haitao Zhang
The propositional connectives. See also init.datatypes and init.logic.
-/
@ -41,6 +41,9 @@ theorem not.intro (H : a → false) : ¬a := H
theorem not_not_intro (Ha : a) : ¬¬a :=
assume Hna : ¬a, absurd Ha Hna
theorem not_imp_not_of_imp {a b : Prop} : (a → b) → ¬b → ¬a :=
assume Pimp Pnb Pa, absurd (Pimp Pa) Pnb
theorem not_not_of_not_implies (H : ¬(a → b)) : ¬¬a :=
assume Hna : ¬a, absurd (assume Ha : a, absurd Ha Hna) H
@ -106,6 +109,21 @@ iff.intro (assume H, and.left H) (assume H, false.elim H)
theorem and_self (a : Prop) : a ∧ a ↔ a :=
iff.intro (assume H, and.left H) (assume H, and.intro H H)
theorem and_imp_eq (a b c : Prop) : (a ∧ b → c) = (a → b → c) :=
propext
(iff.intro (λ Pl a b, Pl (and.intro a b))
(λ Pr Pand, Pr (and.left Pand) (and.right Pand)))
theorem and_eq_right {a b : Prop} (Ha : a) : (a ∧ b) = b :=
propext (iff.intro
(assume Hab, and.elim_right Hab)
(assume Hb, and.intro Ha Hb))
theorem and_eq_left {a b : Prop} (Hb : b) : (a ∧ b) = a :=
propext (iff.intro
(assume Hab, and.elim_left Hab)
(assume Ha, and.intro Ha Hb))
/- or -/
definition not_or (Hna : ¬a) (Hnb : ¬b) : ¬(a b) :=

View file

@ -94,19 +94,30 @@ assume H, by_contradiction (assume Hna : ¬a,
have Hnna : ¬¬a, from not_not_of_not_implies (mt H Hna),
absurd (not_not_elim Hnna) Hna)
theorem forall_not_of_not_exists {A : Type} {P : A → Prop} [D : ∀x, decidable (P x)]
(H : ¬∃x, P x) : ∀x, ¬P x :=
take x, or.elim (em (P x))
(assume Hp : P x, absurd (exists.intro x Hp) H)
(assume Hn : ¬P x, Hn)
theorem forall_not_of_not_exists {A : Type} {p : A → Prop} [D : ∀x, decidable (p x)]
(H : ¬∃x, p x) : ∀x, ¬p x :=
take x, or.elim (em (p x))
(assume Hp : p x, absurd (exists.intro x Hp) H)
(assume Hnp : ¬p x, Hnp)
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 :=
@by_contradiction _ D' (assume H1 : ¬∃x, ¬P x,
have H2 : ∀x, ¬¬P x, from @forall_not_of_not_exists _ _ (take x, decidable_not) H1,
have H3 : ∀x, P x, from take x, @not_not_elim _ (D x) (H2 x),
absurd H3 H)
theorem forall_of_not_exists_not {A : Type} {p : A → Prop} [D : decidable_pred p] :
¬(∃ x, ¬p x) → ∀ x, p x :=
assume Hne, take x, by_contradiction (assume Hnp : ¬ p x, Hne (exists.intro x Hnp))
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 :=
by_contradiction
(assume H1 : ¬∃x, ¬p x,
have H2 : ∀x, ¬¬p x, from forall_not_of_not_exists H1,
have H3 : ∀x, p x, from take x, not_not_elim (H2 x),
absurd H3 H)
theorem exists_of_not_forall_not {A : Type} {p : A → Prop} [D : ∀x, decidable (p x)]
[D' : decidable (∃x, ¬¬p x)] (H : ¬∀x, ¬ p x) :
∃x, p x :=
obtain x (H : ¬¬ p x), from exists_not_of_not_forall H,
exists.intro x (not_not_elim H)
theorem ne_self_iff_false {A : Type} (a : A) : (a ≠ a) ↔ false :=
iff.intro