My solutions to Tao's Analysis I, formalized in Lean
1import Mathlib.Tactic
2import Analysis.Appendix_B_1
3
4/-!
5# Analysis I, Appendix B.2: The decimal representation of real numbers
6
7Am implementation of the decimal representation of Mathlib's real numbers `ℝ`.
8
9This is separate from the way decimal numerals are already represenated in Mathlib. We also represent the integer part of the natural numbers just by `ℕ`, avoiding using the decimal representation from the
10previous section, although we still retain the `Digit` class.
11-/
12
13namespace AppendixB
14
15structure NNRealDecimal where
16 intPart : ℕ
17 fracPart : ℕ → Digit
18
19#check NNRealDecimal.mk
20
21@[coe]
22noncomputable def NNRealDecimal.toNNReal (d:NNRealDecimal) : NNReal :=
23 d.intPart + ∑' i, (d.fracPart i) * (10:NNReal) ^ (-i-1:ℝ)
24
25noncomputable instance NNRealDecimal.instCoeNNReal : Coe NNRealDecimal NNReal where
26 coe := toNNReal
27
28/-- Exercise B.2.1 -/
29theorem NNRealDecimal.toNNReal_conv (d:NNRealDecimal) :
30 Summable fun i ↦ (d.fracPart i) * (10:NNReal) ^ (-i-1:ℝ) := by
31 sorry
32
33theorem NNRealDecimal.surj (x:NNReal) : ∃ d:NNRealDecimal, x = d := by
34 -- This proof is written to follow the structure of the original text.
35 by_cases h : x = 0
36 . use mk 0 fun _ ↦ 0; simp [h, toNNReal]
37 let s : ℕ → ℕ := fun n ↦ ⌊ x * 10^n ⌋₊
38 have hs (n:ℕ) : s n ≤ x * 10^n := Nat.floor_le (by positivity)
39 have hs' (n:ℕ) : x * 10^n < s n + 1 := Nat.lt_floor_add_one _
40 have hdigit (n:ℕ) : ∃ a:Digit, s (n+1) = 10 * s n + (a:ℕ) := by
41 have hl : (10:NNReal) * s n < s (n+1) + 1 := calc
42 _ ≤ 10 * (x * 10^n) := by gcongr; exact hs n
43 _ = x * 10^(n+1) := by ring_nf
44 _ < _ := hs' _
45 have hu : s (n+1) < (10:NNReal) * s n + 10 := calc
46 _ ≤ x * 10^(n+1) := hs (n+1)
47 _ = 10 * (x * 10^n) := by ring_nf
48 _ < 10 * (s n + 1) := by gcongr; exact hs' n
49 _ = _ := by ring
50 norm_cast at hl hu
51 set d := s (n+1) - 10 * s n
52 have hd : d < 10 := by omega
53 have : s (n+1) = 10 * s n + d := by omega
54 use Digit.mk hd
55 choose a ha using hdigit
56 set d := mk (s 0) a; use d
57 have hsum (n:ℕ) : s n * (10:NNReal)^(-n:ℝ) = s 0 + ∑ i ∈ .range n, a i * (10:NNReal)^(-i-1:ℝ) := by
58 induction' n with n hn; simp
59 rw [ha n]; calc
60 _ = s n * (10:NNReal)^(-n:ℝ) + a n * 10^(-n-1:ℝ) := by
61 simp [add_mul]; congr 1 <;> ring_nf
62 rw [mul_assoc, ←NNReal.rpow_add_one (by norm_num)]; congr; ring
63 _ = s 0 + (∑ i ∈ .range n, a i * (10:NNReal)^(-i-1:ℝ) + a n * 10^(-n-1:ℝ)) := by
64 rw [hn]; abel
65 _ = _ := by congr; symm; apply Finset.sum_range_succ
66 have := (d.toNNReal_conv.tendsto_sum_tsum_nat).const_add (s 0:NNReal)
67 convert_to Filter.atTop.Tendsto (fun n ↦ s n * (10:NNReal)^(-n:ℝ)) (nhds (d:NNReal)) at this
68 . ext n; rw [hsum n]
69 apply tendsto_nhds_unique _ this
70 apply Filter.Tendsto.squeeze (g := fun n:ℕ ↦ x - (10:NNReal)^(-n:ℝ)) (h := fun _ ↦ x)
71 . convert Filter.Tendsto.const_sub (c := 0) x _
72 . simp
73 convert NNReal.tendsto_pow_atTop_nhds_zero_of_lt_one
74 (show (1/10:NNReal) < 1 by apply NNReal.div_lt_one_of_lt; norm_num) with n
75 rw [←NNReal.rpow_natCast, one_div, NNReal.inv_rpow, NNReal.rpow_neg]
76 . exact tendsto_const_nhds
77 . intro n; simp; calc
78 _ = (x * 10^n) * (10:NNReal)^(-n:ℝ) := by
79 rw [mul_assoc, ←NNReal.rpow_natCast, ←NNReal.rpow_add (by norm_num)]; simp
80 _ ≤ ((s n:NNReal) + 1)*(10:NNReal)^(-n:ℝ) := by gcongr; exact le_of_lt (hs' n)
81 _ = _ := by ring
82 intro n; simp; calc
83 _ ≤ (x * 10^n) * (10:NNReal)^(-n:ℝ) := by gcongr; exact hs n
84 _ = x := by rw [mul_assoc, ←NNReal.rpow_natCast, ←NNReal.rpow_add (by norm_num)]; simp
85
86/-- Proposition B.2.2 -/
87theorem NNRealDecimal.not_inj : (1:NNReal) = (mk 1 fun _ ↦ 0) ∧ (1:NNReal) = (mk 0 fun _ ↦ 9) := by
88 -- This proof is written to follow the structure of the original text.
89 simp [toNNReal]
90 have := (mk 0 fun _ ↦ 9).toNNReal_conv.tendsto_sum_tsum_nat
91 simp at this
92 apply tendsto_nhds_unique _ this
93 convert_to Filter.atTop.Tendsto (fun n:ℕ ↦ 1 - (10:NNReal)^(-n:ℝ)) (nhds 1) using 2 with n
94 . induction' n with n hn
95 . simp
96 rw [Finset.sum_range_succ, hn, Nat.cast_add, Nat.cast_one, neg_add']
97 have : (10:NNReal)^(-n:ℝ) = 10^(-n-1:ℝ) * 10 := by
98 rw [←NNReal.rpow_add_one (by norm_num)]; simp
99 have hnine : ((9:Digit):ℕ) = 9 := rfl
100 simp [this, hnine, ←NNReal.coe_inj]
101 rw [NNReal.coe_sub, NNReal.coe_sub]
102 . simp; linarith
103 . apply NNReal.rpow_le_one_of_one_le_of_nonpos (by norm_num) (by linarith)
104 rw [←NNReal.rpow_add_one (by norm_num)]
105 apply NNReal.rpow_le_one_of_one_le_of_nonpos (by norm_num) (by linarith)
106 convert Filter.Tendsto.const_sub (f := fun n:ℕ ↦ (10:NNReal)^(-n:ℝ)) (c := 0) 1 _; simp
107 convert NNReal.tendsto_pow_atTop_nhds_zero_of_lt_one
108 (show (1/10:NNReal) < 1 by bound) with n
109 rw [←NNReal.rpow_natCast, one_div, NNReal.inv_rpow, NNReal.rpow_neg]
110
111inductive RealDecimal where
112 | pos : NNRealDecimal → RealDecimal
113 | neg : NNRealDecimal → RealDecimal
114
115noncomputable instance RealDecimal.instCoeReal : Coe RealDecimal ℝ where
116 coe := fun d ↦ match d with
117 | RealDecimal.pos d => d.toNNReal
118 | RealDecimal.neg d => -(d.toNNReal:ℝ)
119
120theorem RealDecimal.surj (x:ℝ) : ∃ d:RealDecimal, x = d := by
121 rcases le_or_gt 0 x with h | h
122 . obtain ⟨ d, hd ⟩ := NNRealDecimal.surj (x.toNNReal); use pos d; simp [←hd, h]
123 . obtain ⟨ d, hd ⟩ := NNRealDecimal.surj ((-x).toNNReal); use neg d; simp [←hd, (show 0 ≤ -x by linarith)]
124
125/-- Exercise B.2.2 -/
126theorem RealDecimal.not_inj_one (d: RealDecimal) : (d:ℝ) = 1 ↔ (d = pos (NNRealDecimal.mk 1 fun _ ↦ 0) ∨ d = pos (NNRealDecimal.mk 0 fun _ ↦ 9)) := by
127 sorry
128
129/-- Exercise B.2.3 -/
130abbrev TerminatingDecimal (x:ℝ) : Prop := ∃ (n:ℤ) (m:ℕ), x = n / (10:ℝ)^m
131
132theorem RealDecimal.not_inj_terminating {x:ℝ} (hx: TerminatingDecimal x) : ∃ d₁ d₂:RealDecimal, d₁ ≠ d₂ ∧ ∀ d: RealDecimal, d = x ↔ d = d₁ ∨ d = d₂ := by sorry
133
134theorem RealDecimal.inj_nonterminating {x:ℝ} (hx: ¬TerminatingDecimal x) : ∃! d:RealDecimal, d = x := by sorry
135
136/-- Exercise B.2.4. This is Corollary 8.3.4, but the intent is to rewrite the proof using the decimal system. -/
137example : Uncountable ℝ := by sorry
138
139
140end AppendixB