bookshelf/Bookshelf/Avigad/Chapter_3.lean

168 lines
4.4 KiB
Plaintext
Raw Normal View History

2023-05-08 19:43:54 +00:00
/-! # Avigad.Chapter3
2023-04-02 14:57:58 +00:00
Propositions and Proofs
-/
/-! ### Exercise 1
2023-05-04 22:37:54 +00:00
Prove the following identities.
-/
2023-05-08 19:43:54 +00:00
namespace Avigad.Chapter3
2023-04-08 16:32:20 +00:00
2023-02-10 21:51:20 +00:00
namespace ex1
open or
2023-02-10 21:51:20 +00:00
variable (p q r : Prop)
-- Commutativity of ∧ and
2023-05-04 22:37:54 +00:00
theorem and_comm' : p ∧ q ↔ q ∧ p :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun ⟨hp, hq⟩ => show q ∧ p from ⟨hq, hp⟩)
(fun ⟨hq, hp⟩ => show p ∧ q from ⟨hp, hq⟩)
2023-05-04 22:37:54 +00:00
theorem or_comm' : p q ↔ q p :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h => h.elim Or.inr Or.inl)
(fun h => h.elim Or.inr Or.inl)
-- Associativity of ∧ and
2023-05-04 22:37:54 +00:00
theorem and_assoc : (p ∧ q) ∧ r ↔ p ∧ (q ∧ r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun ⟨⟨hp, hq⟩, hr⟩ => ⟨hp, hq, hr⟩)
(fun ⟨hp, hq, hr⟩ => ⟨⟨hp, hq⟩, hr⟩)
2023-05-04 22:37:54 +00:00
theorem or_assoc' : (p q) r ↔ p (q r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h₁ => h₁.elim
(fun h₂ => h₂.elim Or.inl (Or.inr ∘ Or.inl))
(Or.inr ∘ Or.inr))
(fun h₁ => h₁.elim
(Or.inl ∘ Or.inl)
(fun h₂ => h₂.elim (Or.inl ∘ Or.inr) Or.inr))
-- Distributivity
2023-05-04 22:37:54 +00:00
theorem and_or_left : p ∧ (q r) ↔ (p ∧ q) (p ∧ r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun ⟨hp, hqr⟩ => hqr.elim (Or.inl ⟨hp, ·⟩) (Or.inr ⟨hp, ·⟩))
(fun h₁ => h₁.elim
(fun ⟨hp, hq⟩ => ⟨hp, Or.inl hq⟩)
(fun ⟨hp, hr⟩ => ⟨hp, Or.inr hr⟩))
2023-05-04 22:37:54 +00:00
theorem or_and_left : p (q ∧ r) ↔ (p q) ∧ (p r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h => h.elim
(fun hp => ⟨Or.inl hp, Or.inl hp⟩)
(fun ⟨hq, hr⟩ => ⟨Or.inr hq, Or.inr hr⟩))
(fun ⟨h₁, h₂⟩ => h₁.elim
Or.inl
(fun hq => h₂.elim Or.inl (fun hr => Or.inr ⟨hq, hr⟩)))
-- Other properties
2023-05-04 22:37:54 +00:00
theorem imp_imp_iff_and_imp : (p → (q → r)) ↔ (p ∧ q → r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h ⟨hp, hq⟩ => h hp hq)
(fun h hp hq => h ⟨hp, hq⟩)
2023-05-04 22:37:54 +00:00
theorem or_imp : ((p q) → r) ↔ (p → r) ∧ (q → r) :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h =>
have h₁ : p → r := h ∘ Or.inl
have h₂ : q → r := h ∘ Or.inr
show (p → r) ∧ (q → r) from ⟨h₁, h₂⟩)
(fun ⟨h₁, h₂⟩ h => h.elim h₁ h₂)
2023-05-04 22:37:54 +00:00
theorem nor_or : ¬(p q) ↔ ¬p ∧ ¬q :=
2023-02-10 21:51:20 +00:00
Iff.intro
(fun h => ⟨h ∘ Or.inl, h ∘ Or.inr⟩)
(fun h₁ h₂ => h₂.elim (absurd · h₁.left) (absurd · h₁.right))
2023-05-04 22:37:54 +00:00
theorem not_and_or_mpr : ¬p ¬q → ¬(p ∧ q) :=
2023-02-10 21:51:20 +00:00
fun h₁ h₂ => h₁.elim (absurd h₂.left ·) (absurd h₂.right ·)
2023-05-04 22:37:54 +00:00
theorem and_not_self : ¬(p ∧ ¬p) :=
2023-02-10 21:51:20 +00:00
fun h => absurd h.left h.right
2023-05-04 22:37:54 +00:00
theorem not_imp_o_and_not : p ∧ ¬q → ¬(p → q) :=
2023-02-10 21:51:20 +00:00
fun ⟨hp, nq⟩ hpq => absurd (hpq hp) nq
2023-05-04 22:37:54 +00:00
theorem false_elim_self : ¬p → (p → q) :=
2023-02-10 21:51:20 +00:00
fun np hp => absurd hp np
2023-05-04 22:37:54 +00:00
theorem not_or_imp_imp : (¬p q) → (p → q) :=
2023-02-10 21:51:20 +00:00
fun npq hp => npq.elim (absurd hp ·) id
2023-05-04 22:37:54 +00:00
theorem or_false_iff : p False ↔ p :=
2023-02-10 21:51:20 +00:00
Iff.intro (fun hpf => hpf.elim id False.elim) Or.inl
2023-05-04 22:37:54 +00:00
theorem and_false_iff : p ∧ False ↔ False :=
2023-02-10 21:51:20 +00:00
Iff.intro (fun ⟨_, hf⟩ => hf) False.elim
2023-05-04 22:37:54 +00:00
theorem imp_imp_not_imp_not : (p → q) → (¬q → ¬p) :=
2023-02-10 21:51:20 +00:00
fun hpq nq hp => absurd (hpq hp) nq
2023-02-10 21:51:20 +00:00
end ex1
/-! ### Exercise 2
2023-05-04 22:37:54 +00:00
Prove the following identities. These require classical reasoning.
-/
2023-04-08 16:32:20 +00:00
2023-02-10 21:51:20 +00:00
namespace ex2
2023-02-10 21:51:20 +00:00
open Classical
2023-02-10 21:51:20 +00:00
variable (p q r s : Prop)
2023-05-04 22:37:54 +00:00
theorem imp_or_mp (hp : p) : (p → r s) → ((p → r) (p → s)) :=
2023-02-10 21:51:20 +00:00
fun h => (h hp).elim
(fun hr => Or.inl (fun _ => hr))
(fun hs => Or.inr (fun _ => hs))
2023-05-04 22:37:54 +00:00
theorem not_and_iff_or_not : ¬(p ∧ q) → ¬p ¬q :=
2023-02-10 21:51:20 +00:00
fun npq => (em p).elim
(fun hp => (em q).elim
(fun hq => False.elim (npq ⟨hp, hq⟩))
Or.inr)
Or.inl
2023-05-04 22:37:54 +00:00
theorem not_imp_mp : ¬(p → q) → p ∧ ¬q :=
2023-02-10 21:51:20 +00:00
fun h =>
have lhs : p := byContradiction
fun np => h (fun (hp : p) => absurd hp np)
⟨lhs, fun hq => h (fun _ => hq)⟩
2023-05-04 22:37:54 +00:00
theorem not_or_of_imp : (p → q) → (¬p q) :=
2023-02-10 21:51:20 +00:00
fun hpq => (em p).elim (fun hp => Or.inr (hpq hp)) Or.inl
2023-05-04 22:37:54 +00:00
theorem not_imp_not_imp_imp : (¬q → ¬p) → (p → q) :=
2023-02-10 21:51:20 +00:00
fun h hp => byContradiction
fun nq => absurd hp (h nq)
2023-05-04 22:37:54 +00:00
theorem or_not : p ¬p := em p
2023-05-04 22:37:54 +00:00
theorem imp_imp_imp : (((p → q) → p) → p) :=
2023-02-10 21:51:20 +00:00
fun h => byContradiction
fun np =>
suffices hp : p from absurd hp np
h (fun (hp : p) => absurd hp np)
2023-02-10 21:51:20 +00:00
end ex2
/-! ### Exercise 3
2023-05-04 22:37:54 +00:00
Prove `¬(p ↔ ¬p)` without using classical logic.
-/
2023-04-08 16:32:20 +00:00
2023-02-10 21:51:20 +00:00
namespace ex3
2023-02-10 21:51:20 +00:00
variable (p : Prop)
2023-05-04 22:37:54 +00:00
theorem iff_not_self (hp : p) : ¬(p ↔ ¬p) :=
2023-02-10 21:51:20 +00:00
fun h => absurd hp (Iff.mp h hp)
2023-02-10 21:51:20 +00:00
end ex3
2023-05-04 22:37:54 +00:00
end Avigad.Chapter3