2023-08-24 00:23:28 +00:00
|
|
|
|
import Mathlib.Data.Finset.Basic
|
|
|
|
|
import Mathlib.Data.Set.Finite
|
2023-08-16 18:46:16 +00:00
|
|
|
|
import Mathlib.Data.Set.Function
|
|
|
|
|
import Mathlib.Data.Rel
|
|
|
|
|
|
|
|
|
|
/-! # Enderton.Set.Chapter_6
|
|
|
|
|
|
|
|
|
|
Cardinal Numbers and the Axiom of Choice
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
namespace Enderton.Set.Chapter_6
|
|
|
|
|
|
|
|
|
|
/-! #### Theorem 6A
|
|
|
|
|
|
|
|
|
|
For any sets `A`, `B`, and `C`,
|
|
|
|
|
|
|
|
|
|
(a) `A ≈ A`.
|
|
|
|
|
(b) If `A ≈ B`, then `B ≈ A`.
|
|
|
|
|
(c) If `A ≈ B` and `B ≈ C`, then `A ≈ C`.
|
|
|
|
|
-/
|
|
|
|
|
|
|
|
|
|
theorem theorem_6a_a (A : Set α)
|
2023-08-17 18:56:39 +00:00
|
|
|
|
: ∃ F, Set.BijOn F A A := by
|
2023-08-16 18:46:16 +00:00
|
|
|
|
refine ⟨fun x => x, ?_⟩
|
|
|
|
|
unfold Set.BijOn Set.MapsTo Set.InjOn Set.SurjOn
|
|
|
|
|
simp only [imp_self, implies_true, Set.image_id', true_and]
|
|
|
|
|
exact Eq.subset rfl
|
|
|
|
|
|
|
|
|
|
theorem theorem_6a_b [Nonempty α] (A : Set α) (B : Set β)
|
2023-08-17 18:56:39 +00:00
|
|
|
|
(F : α → β) (hF : Set.BijOn F A B)
|
|
|
|
|
: ∃ G, Set.BijOn G B A := by
|
|
|
|
|
refine ⟨Function.invFunOn F A, ?_⟩
|
|
|
|
|
exact (Set.bijOn_comm $ Set.BijOn.invOn_invFunOn hF).mpr hF
|
2023-08-16 18:46:16 +00:00
|
|
|
|
|
|
|
|
|
theorem theorem_6a_c (A : Set α) (B : Set β) (C : Set γ)
|
2023-08-17 18:56:39 +00:00
|
|
|
|
(F : α → β) (hF : Set.BijOn F A B)
|
|
|
|
|
(G : β → γ) (hG : Set.BijOn G B C)
|
|
|
|
|
: ∃ H, Set.BijOn H A C := by
|
|
|
|
|
exact ⟨G ∘ F, Set.BijOn.comp hG hF⟩
|
2023-08-16 18:46:16 +00:00
|
|
|
|
|
2023-08-24 00:23:28 +00:00
|
|
|
|
/-- #### Theorem 6B
|
|
|
|
|
|
|
|
|
|
No set is equinumerous to its powerset.
|
|
|
|
|
-/
|
|
|
|
|
theorem theorem_6b (A : Set α)
|
|
|
|
|
: ∀ f, ¬ Set.BijOn f A (𝒫 A) := by
|
|
|
|
|
intro f hf
|
|
|
|
|
unfold Set.BijOn at hf
|
|
|
|
|
let φ := { a ∈ A | a ∉ f a }
|
|
|
|
|
suffices ∀ a ∈ A, f a ≠ φ by
|
|
|
|
|
have hφ := hf.right.right (show φ ∈ 𝒫 A by simp)
|
|
|
|
|
have ⟨a, ha⟩ := hφ
|
|
|
|
|
exact absurd ha.right (this a ha.left)
|
|
|
|
|
intro a ha hfa
|
|
|
|
|
by_cases h : a ∈ f a
|
|
|
|
|
· have h' := h
|
|
|
|
|
rw [hfa] at h
|
|
|
|
|
simp only [Set.mem_setOf_eq] at h
|
|
|
|
|
exact absurd h' h.right
|
|
|
|
|
· rw [Set.Subset.antisymm_iff] at hfa
|
|
|
|
|
have := hfa.right ⟨ha, h⟩
|
|
|
|
|
exact absurd this h
|
|
|
|
|
|
|
|
|
|
/-- #### Pigeonhole Principle
|
|
|
|
|
|
|
|
|
|
No natural number is equinumerous to a proper subset of itself.
|
|
|
|
|
-/
|
|
|
|
|
theorem pigeonhole_principle (m n : ℕ) (hm : m < n)
|
|
|
|
|
: ∀ f : Fin m → Fin n, ¬ Function.Bijective f := by
|
2023-08-24 13:50:47 +00:00
|
|
|
|
induction n with
|
|
|
|
|
| zero =>
|
|
|
|
|
intro f hf
|
|
|
|
|
simp at hm
|
|
|
|
|
| succ n ih =>
|
|
|
|
|
sorry
|
2023-08-24 00:23:28 +00:00
|
|
|
|
|
|
|
|
|
/-- #### Corollary 6C
|
|
|
|
|
|
|
|
|
|
No finite set is equinumerous to a proper subset of itself.
|
|
|
|
|
-/
|
|
|
|
|
theorem corollary_6c (S S' : Finset α) (hS : S' ⊂ S)
|
|
|
|
|
: ∀ f : S → S', ¬ Function.Bijective f := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Corollary 6D (a)
|
|
|
|
|
|
|
|
|
|
Any set equinumerous to a proper subset of itself is infinite.
|
|
|
|
|
-/
|
|
|
|
|
theorem corollary_6d_a (S S' : Set α) (hS : S' ⊂ S) (hf : S' ≃ S)
|
|
|
|
|
: Set.Infinite S := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Corollary 6D (b)
|
|
|
|
|
|
|
|
|
|
The set `ω` is infinite.
|
|
|
|
|
-/
|
|
|
|
|
theorem corollary_6d_b
|
|
|
|
|
: Set.Infinite (@Set.univ ℕ) := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Corollary 6E
|
|
|
|
|
|
|
|
|
|
Any finite set is equinumerous to a unique natural number.
|
|
|
|
|
-/
|
2023-08-24 13:50:47 +00:00
|
|
|
|
theorem corollary_6e (S : Set α) (hn : S ≃ Fin n) (hm : S ≃ Fin m)
|
|
|
|
|
: m = n := by
|
2023-08-24 00:23:28 +00:00
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Lemma 6F
|
|
|
|
|
|
|
|
|
|
If `C` is a proper subset of a natural number `n`, then `C ≈ m` for some `m`
|
|
|
|
|
less than `n`.
|
|
|
|
|
-/
|
2023-08-24 13:50:47 +00:00
|
|
|
|
lemma lemma_6f {n : ℕ} (hC : C ⊂ Finset.range n)
|
2023-08-24 00:23:28 +00:00
|
|
|
|
: ∃ m : ℕ, m < n ∧ ∃ f : C → Fin m, Function.Bijective f := by
|
|
|
|
|
sorry
|
|
|
|
|
|
2023-08-24 13:50:47 +00:00
|
|
|
|
theorem corollary_6g (S S' : Set α) (hS : Finite S) (hS' : S' ⊆ S)
|
|
|
|
|
: Finite S' := by
|
|
|
|
|
sorry
|
|
|
|
|
|
2023-08-17 20:10:21 +00:00
|
|
|
|
/-- #### Exercise 6.1
|
|
|
|
|
|
|
|
|
|
Show that the equation
|
|
|
|
|
```
|
|
|
|
|
f(m, n) = 2ᵐ(2n + 1) - 1
|
|
|
|
|
```
|
|
|
|
|
defines a one-to-one correspondence between `ω × ω` and `ω`.
|
|
|
|
|
-/
|
|
|
|
|
theorem exercise_6_1
|
|
|
|
|
: Function.Bijective (fun p : ℕ × ℕ => 2 ^ p.1 * (2 * p.2 + 1) - 1) := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Exercise 6.2
|
|
|
|
|
|
|
|
|
|
Show that in Fig. 32 we have:
|
|
|
|
|
```
|
|
|
|
|
J(m, n) = [1 + 2 + ⋯ + (m + n)] + m
|
|
|
|
|
= (1 / 2)[(m + n)² + 3m + n].
|
|
|
|
|
```
|
|
|
|
|
-/
|
|
|
|
|
theorem exercise_6_2
|
|
|
|
|
: Function.Bijective
|
|
|
|
|
(fun p : ℕ × ℕ => (1 / 2) * ((p.1 + p.2) ^ 2 + 3 * p.1 + p.2)) := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Exercise 6.3
|
|
|
|
|
|
|
|
|
|
Find a one-to-one correspondence between the open unit interval `(0, 1)` and `ℝ`
|
|
|
|
|
that takes rationals to rationals and irrationals to irrationals.
|
|
|
|
|
-/
|
|
|
|
|
theorem exercise_6_3
|
|
|
|
|
: True := by
|
|
|
|
|
sorry
|
|
|
|
|
|
|
|
|
|
/-- #### Exercise 6.4
|
|
|
|
|
|
|
|
|
|
Construct a one-to-one correspondence between the closed unit interval
|
|
|
|
|
```
|
|
|
|
|
[0, 1] = {x ∈ ℝ | 0 ≤ x ≤ 1}
|
|
|
|
|
```
|
|
|
|
|
and the open unit interval `(0, 1)`.
|
|
|
|
|
-/
|
|
|
|
|
theorem exercise_6_4
|
|
|
|
|
: ∃ F, Set.BijOn F (Set.Ioo 0 1) (@Set.univ ℝ) := by
|
|
|
|
|
sorry
|
|
|
|
|
|
2023-08-16 18:46:16 +00:00
|
|
|
|
end Enderton.Set.Chapter_6
|