├── .gitignore ├── ComputableReal.lean ├── ComputableReal ├── ComputableRSeq.lean ├── ComputableReal.lean ├── IsComputable.lean ├── IsComputableC.lean ├── SpecialFunctions.lean ├── SpecialFunctions │ ├── Basic.lean │ ├── Exp.lean │ ├── Pi.lean │ └── Sqrt.lean ├── aux_lemmas.lean └── examples.lean ├── LICENSE ├── README.md ├── lake-manifest.json ├── lakefile.lean └── lean-toolchain /.gitignore: -------------------------------------------------------------------------------- 1 | /build 2 | /lake-packages/* 3 | .lake/* 4 | -------------------------------------------------------------------------------- /ComputableReal.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.ComputableRSeq 2 | import ComputableReal.ComputableReal 3 | import ComputableReal.IsComputable 4 | import ComputableReal.IsComputableC 5 | import ComputableReal.SpecialFunctions 6 | -------------------------------------------------------------------------------- /ComputableReal/ComputableRSeq.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Algebra.Order.Interval.Basic 2 | import Mathlib.Data.Real.Archimedean 3 | import Mathlib.Data.Sign 4 | import Mathlib.Tactic.Rify 5 | 6 | import ComputableReal.aux_lemmas 7 | 8 | namespace QInterval 9 | 10 | scoped notation "ℚInterval" => NonemptyInterval ℚ 11 | 12 | scoped instance (priority := 100) instMemℝℚInterval : Membership ℝ ℚInterval := 13 | ⟨fun s a => s.fst ≤ a ∧ a ≤ s.snd⟩ 14 | 15 | section mul 16 | /--Multiplication on intervals of ℚ. TODO: Should generalize to any LinearOrderedField... -/ 17 | def mul_pair (x y : ℚInterval) : ℚInterval := 18 | let ⟨⟨xl,xu⟩,_⟩ := x 19 | let ⟨⟨yl,yu⟩,_⟩ := y 20 | ⟨⟨min (min (xl*yl) (xu*yl)) (min (xl*yu) (xu*yu)), 21 | max (max (xl*yl) (xu*yl)) (max (xl*yu) (xu*yu))⟩, 22 | by simp only [le_max_iff, min_le_iff, le_refl, true_or, or_true, or_self]⟩ 23 | 24 | /--Multiplication of intervals by a ℚ. TODO: Should generalize to any LinearOrderedField -/ 25 | def mul_q (x : ℚInterval) (y : ℚ) : ℚInterval := 26 | if h : y ≥ 0 then 27 | ⟨⟨x.fst * y, x.snd * y⟩, by dsimp; nlinarith [x.2]⟩ 28 | else 29 | ⟨⟨x.snd * y, x.fst * y⟩, by dsimp; nlinarith [x.2]⟩ 30 | 31 | scoped instance : Mul (ℚInterval) := 32 | ⟨mul_pair⟩ 33 | 34 | scoped instance : HMul (ℚInterval) ℚ (ℚInterval) := 35 | ⟨mul_q⟩ 36 | 37 | scoped instance : HDiv (ℚInterval) ℚ (ℚInterval) := 38 | ⟨fun x y ↦ x * y⁻¹⟩ 39 | 40 | section slow 41 | set_option maxHeartbeats 400000 42 | theorem mul_pair_lb_is_lb {x y : ℚInterval} : ∀ xv ∈ x, ∀ yv ∈ y, 43 | (mul_pair x y).fst ≤ xv * yv := by 44 | intro xv ⟨hxl,hxu⟩ yv ⟨hyl,hyu⟩ 45 | dsimp [mul_pair] 46 | push_cast 47 | rcases le_or_lt xv 0 with hxn|hxp 48 | all_goals rcases le_or_lt (y.fst:ℝ) 0 with hyln|hylp 49 | all_goals rcases le_or_lt (y.snd:ℝ) 0 with hyun|hyup 50 | all_goals try linarith 51 | all_goals repeat rw [min_def] 52 | all_goals split_ifs with h₁ h₂ h₃ h₃ h₂ h₃ h₃ 53 | all_goals try nlinarith 54 | 55 | theorem mul_pair_ub_is_ub {x y : ℚInterval} : ∀ xv ∈ x, ∀ yv ∈ y, 56 | (mul_pair x y).snd ≥ xv * yv := by 57 | intro xv ⟨hxl,hxu⟩ yv ⟨hyl,hyu⟩ 58 | dsimp [mul_pair] 59 | push_cast 60 | rcases le_or_lt xv 0 with hxn|hxp 61 | all_goals rcases le_or_lt (y.1.1:ℝ) 0 with hyln|hylp 62 | all_goals rcases le_or_lt (y.1.2:ℝ) 0 with hyun|hyup 63 | all_goals try linarith 64 | all_goals repeat rw [max_def] 65 | all_goals split_ifs with h₁ h₂ h₃ h₃ h₂ h₃ h₃ 66 | all_goals try nlinarith 67 | 68 | end slow 69 | 70 | theorem mem_mul_pair {x y : ℚInterval} : ∀ xv ∈ x, ∀ yv ∈ y, xv * yv ∈ mul_pair x y := 71 | fun _ hx _ hy ↦ ⟨mul_pair_lb_is_lb _ hx _ hy, mul_pair_ub_is_ub _ hx _ hy⟩ 72 | 73 | end mul 74 | 75 | scoped instance instRatCastQInterval : RatCast ℚInterval := 76 | ⟨fun q ↦ NonemptyInterval.pure q⟩ 77 | 78 | scoped instance instOfNatQInterval : OfNat ℚInterval n := 79 | ⟨NonemptyInterval.pure n⟩ 80 | 81 | end QInterval 82 | 83 | /-- Type class for sequences that converge to some real number from above and below. `lub` is a 84 | function that gives upper *and* lower bounds, bundled so it can reuse computation. `hcau` 85 | asserts that the two bounds are Cauchy sequences, `hlub` asserts that they're valid 86 | lower and upper bounds, and `heq'` asserts that they converge to a common value. Use 87 | `ComputableℝSeq.mk` to construct with regards to a reference real value. 88 | 89 | The defs `lb`, `ub` are the actual CauSeq's , `val` is the associated real number, 90 | and `hlb`, `hub`, and `heq` relate `lb` `ub` and `val` to each other. -/ 91 | structure ComputableℝSeq where 92 | mk' :: 93 | lub : ℕ → NonemptyInterval ℚ 94 | hcl : IsCauSeq abs fun n ↦ (lub n).fst 95 | hcu : IsCauSeq abs fun n ↦ (lub n).snd 96 | hlub : ∀n, (lub n).fst ≤ (Real.mk ⟨fun n ↦ (lub n).fst, hcl⟩) ∧ 97 | (lub n).snd ≥ (Real.mk ⟨fun n ↦ (lub n).fst, hcl⟩) 98 | heq' : let lb : CauSeq ℚ abs := ⟨fun n ↦ (lub n).fst, hcl⟩ 99 | let ub : CauSeq ℚ abs := ⟨fun n ↦ (lub n).snd, hcu⟩ 100 | lb ≈ ub 101 | 102 | namespace ComputableℝSeq 103 | 104 | open scoped QInterval 105 | 106 | def lb (x : ComputableℝSeq) : CauSeq ℚ abs := ⟨fun n ↦ (x.lub n).fst, x.hcl⟩ 107 | 108 | def ub (x : ComputableℝSeq) : CauSeq ℚ abs := ⟨fun n ↦ (x.lub n).snd, x.hcu⟩ 109 | 110 | /-- Get the real value determined by the sequence. (Irreducibly) given here as the limit of 111 | the lower bound sequence. -/ 112 | irreducible_def val (x : ComputableℝSeq) : ℝ := Real.mk x.lb 113 | 114 | theorem heq (x : ComputableℝSeq) : x.lb ≈ x.ub := 115 | x.heq' 116 | 117 | theorem lb_eq_ub (x : ComputableℝSeq) : Real.mk x.lb = Real.mk x.ub := 118 | Real.mk_eq.2 x.heq' 119 | 120 | theorem val_eq_mk_lb (x : ComputableℝSeq) : x.val = Real.mk x.lb := 121 | val_def x 122 | 123 | theorem val_eq_mk_ub (x : ComputableℝSeq) : x.val = Real.mk x.ub := 124 | x.val_eq_mk_lb.trans x.lb_eq_ub 125 | 126 | theorem hlb (x : ComputableℝSeq) : ∀n, x.lb n ≤ x.val := 127 | fun n ↦ val_eq_mk_lb _ ▸ (x.hlub n).1 128 | 129 | theorem hub (x : ComputableℝSeq) : ∀n, x.ub n ≥ x.val := 130 | fun n ↦ val_eq_mk_lb _ ▸ (x.hlub n).2 131 | 132 | theorem val_mem_interval (x : ComputableℝSeq) : ∀n, x.val ∈ x.lub n := 133 | fun n ↦ ⟨x.hlb n, x.hub n⟩ 134 | 135 | private theorem val_uniq' {x : ℝ} {lb ub : CauSeq ℚ abs} (hlb : ∀n, lb n ≤ x) 136 | (hub : ∀n, ub n ≥ x) (heq : lb ≈ ub) : Real.mk lb = x := 137 | (Real.of_near lb x (fun εℝ hεℝ ↦ 138 | let ⟨ε, ⟨hε₁, hε₂⟩⟩ := exists_rat_btwn hεℝ 139 | let ⟨i,hi⟩ := heq ε (Rat.cast_pos.1 hε₁) 140 | ⟨i, fun j hj ↦ by 141 | replace hi := hi j hj 142 | have hl₁ := hlb j 143 | have hu₂ := hub j 144 | rify at hi hl₁ hu₂ hε₁ 145 | rw [abs_ite_le] at hi ⊢ 146 | split_ifs at hi ⊢ 147 | <;> linarith⟩)).2 148 | 149 | /-- If a real number x is bounded below and above by a sequence, it must be the value of that sequence. -/ 150 | theorem val_uniq {x : ℝ} {s : ComputableℝSeq} (hlb : ∀n, s.lb n ≤ x) (hub : ∀n, s.ub n ≥ x) : 151 | s.val = x := 152 | s.val_def ▸ val_uniq' hlb hub s.heq 153 | 154 | /-- Make a computable sequence for x from a separate lower and upper bound CauSeq. -/ 155 | def mk (x : ℝ) (lub : ℕ → ℚInterval) 156 | (hcl : IsCauSeq abs (fun n ↦ (lub n).fst)) 157 | (hcu : IsCauSeq abs (fun n ↦ (lub n).snd)) 158 | (hlb : ∀n, (lub n).fst ≤ x) 159 | (hub : ∀n, (lub n).snd ≥ x) 160 | (heq : let lb : CauSeq ℚ abs := ⟨fun n ↦ (lub n).fst, hcl⟩ 161 | let ub : CauSeq ℚ abs := ⟨fun n ↦ (lub n).snd, hcu⟩ 162 | lb ≈ ub) : ComputableℝSeq where 163 | lub := lub 164 | hcl := hcl 165 | hcu := hcu 166 | heq' := heq 167 | hlub n := by 168 | rw [val_uniq' hlb hub heq] 169 | exact ⟨hlb n, hub n⟩ 170 | 171 | theorem mk_val_eq_val : (mk x v h₁ h₂ h₃ h₄ h₅).val = x := 172 | val_uniq (by convert h₃) (by convert h₄) 173 | 174 | theorem lb_le_ub (x : ComputableℝSeq) : ∀n, x.lb n ≤ x.ub n := 175 | fun n ↦ Rat.cast_le.mp (le_trans (x.hlb n) (x.hub n)) 176 | 177 | @[ext] 178 | theorem ext {x y : ComputableℝSeq} (h₁ : ∀ n, x.lb n = y.lb n) (h₂ : ∀ n, x.ub n = y.ub n) : x = y := 179 | mk'.injEq _ _ _ _ _ _ _ _ _ _ ▸ (funext fun n ↦ NonemptyInterval.ext (Prod.ext (h₁ n) (h₂ n))) 180 | 181 | /-- All rational numbers `q` have a computable sequence: the constant sequence `q`. -/ 182 | def ofRat (q : ℚ) : ComputableℝSeq := 183 | mk q 184 | (fun _ ↦ NonemptyInterval.pure q) 185 | (IsCauSeq.const q) (IsCauSeq.const q) 186 | (fun _ ↦ rfl.le) (fun _ ↦ rfl.le) 187 | (Real.mk_eq.mp rfl) 188 | 189 | instance natCast : NatCast ComputableℝSeq where natCast n := ofRat n 190 | 191 | instance intCast : IntCast ComputableℝSeq where intCast z := ofRat z 192 | 193 | instance ratCast : RatCast ComputableℝSeq where ratCast q := ofRat q 194 | 195 | def add (x : ComputableℝSeq) (y : ComputableℝSeq) : ComputableℝSeq := 196 | mk (x.val + y.val) 197 | (fun n ↦ x.lub n + y.lub n) 198 | (IsCauSeq.add x.hcl y.hcl) 199 | (IsCauSeq.add x.hcu y.hcu) 200 | (by 201 | intro n 202 | rw [NonemptyInterval.fst_add] 203 | push_cast 204 | exact add_le_add (x.hlb n) (y.hlb n)) 205 | (by 206 | intro n 207 | rw [NonemptyInterval.snd_add] 208 | push_cast 209 | exact add_le_add (x.hub n) (y.hub n)) 210 | (have := CauSeq.add_equiv_add x.heq y.heq; this) --TODO why does 'inlining' the have not work 211 | 212 | def neg (x : ComputableℝSeq) : ComputableℝSeq := 213 | mk (-x.val) 214 | (fun n ↦ -x.lub n) 215 | (IsCauSeq.neg x.hcu) 216 | (IsCauSeq.neg x.hcl) 217 | (by simpa using x.hub ·) 218 | (by simpa using x.hlb ·) 219 | (have := CauSeq.neg_equiv_neg (Setoid.symm x.heq); this) 220 | 221 | def sub (x : ComputableℝSeq) (y : ComputableℝSeq) : ComputableℝSeq := 222 | add x (neg y) 223 | 224 | -- def applyℚMono (x : ComputableℝSeq) (fq : ℚ → ℚ) (fr : ℝ → ℝ) (hfr : ∀q, fq q = fr q) 225 | -- (hf₁ : Monotone fr) (hf₂ : Continuous fr) : ComputableℝSeq := 226 | -- mk (x := fr x.val) 227 | -- (lb := x.lb.apply fq (hf₂.continuous_embed fq hfr)) 228 | -- (ub := x.ub.apply fq (hf₂.continuous_embed fq hfr)) 229 | -- (hlb := fun n ↦ (hfr _).symm ▸ (hf₁ (x.hlb n))) 230 | -- (hub := fun n ↦ (hfr _).symm ▸ (hf₁ (x.hub n))) 231 | -- (heq := CauSeq.equiv_apply x.lb x.ub fq (hf₂.continuous_embed fq hfr) x.heq) 232 | 233 | -- @[simp] 234 | -- theorem val_applyℚMono : (applyℚMono x fq fr h₁ h₂ h₃).val = fr x.val := 235 | -- mk_val_eq_val 236 | 237 | -- def applyℚAnti (x : ComputableℝSeq) (fq : ℚ → ℚ) (fr : ℝ → ℝ) (hfr : ∀q, fq q = fr q) 238 | -- (hf₁ : Antitone fr) (hf₂ : Continuous fr) : ComputableℝSeq := 239 | -- applyℚMono (neg x) (fq∘Neg.neg) (fr∘Neg.neg) 240 | -- (fun q ↦ by have := hfr (-q); rwa [Rat.cast_neg] at this) 241 | -- (hf₁.comp monotone_id.neg) 242 | -- (hf₂.comp ContinuousNeg.continuous_neg) 243 | 244 | -- @[simp] 245 | -- theorem val_applyℚAnti : (applyℚAnti x fq fr h₁ h₂ h₃).val = fr x.val := by 246 | -- rw [applyℚAnti, val_applyℚMono, neg, mk_val_eq_val] 247 | -- dsimp 248 | -- rw [neg_neg] 249 | 250 | -- --Faster one for rational multiplcation 251 | -- def lb_mul_q [hx : ComputableℝSeq x] : CauSeq ℚ qabs := 252 | -- if q ≥ 0 then hx.lb * CauSeq.const qabs q else hx.ub * CauSeq.const qabs q 253 | 254 | -- def ub_mul_q [hx : ComputableℝSeq x] : CauSeq ℚ qabs := 255 | -- if q ≥ 0 then hx.ub * CauSeq.const qabs q else hx.lb * CauSeq.const qabs q 256 | 257 | -- /- Multiplication of two computable sequences. Can't just use CauSeq mul because that 258 | -- no longer gives correct upper/lower bounds. -/ 259 | -- def ComputableℝSeqMul [hx : ComputableℝSeq x] : ComputableℝSeq (x * q) where 260 | -- lb := lb_mul_q x q 261 | -- ub := ub_mul_q x q 262 | -- hlb n := by 263 | -- simp_rw [lb_mul_q, min_def] 264 | -- by_cases hq : (q:ℝ) > 0 265 | -- <;> split_ifs with h 266 | -- <;> rify at * 267 | -- <;> nlinarith (config := {splitNe := true}) [hx.hlb n, hx.hub n] 268 | -- hub n := by 269 | -- simp_rw [ub_mul_q, max_def] 270 | -- by_cases hq : (q:ℝ) > 0 271 | -- <;> split_ifs with h 272 | -- <;> rify at * 273 | -- <;> nlinarith (config := {splitNe := true}) [hx.hlb n, hx.hub n] 274 | -- heq := by 275 | -- have : (ub_mul_q x q - lb_mul_q x q) 276 | -- = fun n => (abs (ub x n - lb x n)) * (abs q) := by 277 | -- funext n 278 | -- dsimp 279 | -- simp_rw [ub_mul_q, lb_mul_q] 280 | -- simp_rw [min_def, max_def, abs_ite_le] 281 | -- split_ifs <;> nlinarith 282 | -- rw [this] 283 | -- apply IsCauSeq.mul 284 | -- · intro ε hε 285 | -- obtain ⟨i, hi⟩ := hx.hgap ε hε 286 | -- use i 287 | -- intro j hj 288 | -- replace hi := hi j hj 289 | -- simp_rw [abs_ite_le] at hi ⊢ 290 | -- split_ifs at hi ⊢ 291 | -- <;> dsimp at * <;> linarith 292 | -- · exact IsCauSeq.const _ 293 | 294 | -- instance instComputableQMul [ComputableℝSeq x] : ComputableℝSeq (q * x) := 295 | -- mul_comm x q ▸ instComputableMulQ x q 296 | 297 | 298 | /-- "Bundled" multiplication to give lower and upper bounds. This bundling avoids the need to 299 | call lb and ub separately for each half (which, in a large product, leads to an exponential 300 | slowdown). This could be further optimized to use only two ℚ multiplications instead of four, 301 | when the sign is apparent. -/ 302 | def mul' (x : ComputableℝSeq) (y : ComputableℝSeq) : ℕ → ℚInterval := 303 | fun n ↦ QInterval.mul_pair (x.lub n) (y.lub n) 304 | 305 | /-- More friendly expression for the lower bound for multiplication, as a CauSeq. -/ 306 | def mul_lb (x : ComputableℝSeq) (y : ComputableℝSeq) : CauSeq ℚ abs := 307 | ((x.lb * y.lb) ⊓ (x.ub * y.lb)) ⊓ ((x.lb * y.ub) ⊓ (x.ub * y.ub)) 308 | 309 | /-- More friendly expression for the lower bound for multiplication, as a CauSeq. -/ 310 | def mul_ub (x : ComputableℝSeq) (y : ComputableℝSeq) : CauSeq ℚ abs := 311 | ((x.lb * y.lb) ⊔ (x.ub * y.lb)) ⊔ ((x.lb * y.ub) ⊔ (x.ub * y.ub)) 312 | 313 | /-- The lower bounds from `mul'` are precisely the same sequence as `mul_lb`. -/ 314 | theorem fst_mul'_eq_mul_lb : (fun i ↦ i.fst) ∘ mul' x y = (mul_lb x y).1 := by 315 | ext n 316 | dsimp 317 | rw [mul', mul_lb] 318 | congr 319 | 320 | /-- The upper bounds from `mul'` are precisely the same sequence as `mul_ub`. -/ 321 | theorem snd_mul'_eq_mul_ub : (fun i ↦ i.snd) ∘ mul' x y = (mul_ub x y).1 := by 322 | ext n 323 | dsimp 324 | rw [mul', mul_ub] 325 | congr 326 | 327 | /-- The lower bounds from `mul'` form a Cauchy sequence. -/ 328 | theorem mul'_fst_iscau : IsCauSeq abs ((fun i ↦ i.fst) ∘ (mul' x y)) := 329 | fst_mul'_eq_mul_lb ▸ Subtype.property _ 330 | 331 | /-- The upper bounds from `mul'` form a Cauchy sequence. -/ 332 | theorem mul'_snd_iscau : IsCauSeq abs ((fun i ↦ i.snd) ∘ (mul' x y)) := 333 | snd_mul'_eq_mul_ub ▸ Subtype.property _ 334 | 335 | theorem lb_ub_mul_equiv (x : ComputableℝSeq) (y : ComputableℝSeq) : 336 | mul_lb x y ≈ mul_ub x y := by 337 | have : x.lb ≈ x.lb := by rfl 338 | have : x.ub ≈ x.ub := by rfl 339 | have : y.lb ≈ y.lb := by rfl 340 | have : y.ub ≈ y.ub := by rfl 341 | have := x.heq 342 | have := Setoid.symm x.heq 343 | have := y.heq 344 | have := Setoid.symm y.heq 345 | dsimp [mul_lb, mul_ub] 346 | apply CauSeq.inf_equiv_of_equivs 347 | <;> apply CauSeq.inf_equiv_of_equivs 348 | <;> apply CauSeq.equiv_sup_of_equivs 349 | <;> apply CauSeq.equiv_sup_of_equivs 350 | <;> exact CauSeq.mul_equiv_mul ‹_› ‹_› 351 | 352 | theorem mul_lb_is_lb (x : ComputableℝSeq) (y : ComputableℝSeq) (n : ℕ) : 353 | (mul_lb x y).1 n ≤ x.val * y.val := 354 | QInterval.mul_pair_lb_is_lb _ (x.val_mem_interval n) _ (y.val_mem_interval n) 355 | 356 | theorem mul_ub_is_ub (x : ComputableℝSeq) (y : ComputableℝSeq) (n : ℕ) : 357 | (mul_ub x y).1 n ≥ x.val * y.val := 358 | QInterval.mul_pair_ub_is_ub _ (x.val_mem_interval n) _ (y.val_mem_interval n) 359 | 360 | def mul (x : ComputableℝSeq) (y : ComputableℝSeq) : ComputableℝSeq where 361 | lub := mul' x y 362 | hcl := mul'_fst_iscau 363 | hcu := mul'_snd_iscau 364 | heq' := by convert lb_ub_mul_equiv x y 365 | hlub n := 366 | let h₀ : Real.mk _ = x.val * y.val := by 367 | apply val_uniq' (mul_lb_is_lb x y) (mul_ub_is_ub x y) 368 | convert lb_ub_mul_equiv x y 369 | h₀ ▸ QInterval.mem_mul_pair _ (x.val_mem_interval n) _ (y.val_mem_interval n) 370 | 371 | instance instComputableZero : Zero ComputableℝSeq := 372 | ⟨(0 : ℕ)⟩ 373 | 374 | instance instComputableOne : One ComputableℝSeq := 375 | ⟨(1 : ℕ)⟩ 376 | 377 | instance instAdd : Add ComputableℝSeq := 378 | ⟨add⟩ 379 | 380 | instance instNeg : Neg ComputableℝSeq := 381 | ⟨neg⟩ 382 | 383 | instance instSub : Sub ComputableℝSeq := 384 | ⟨sub⟩ 385 | 386 | instance instMul : Mul ComputableℝSeq := 387 | ⟨mul⟩ 388 | 389 | instance instInh : Inhabited ComputableℝSeq := 390 | ⟨0⟩ 391 | 392 | section simps 393 | 394 | variable (x y : ComputableℝSeq) 395 | 396 | @[simp] 397 | theorem natCast_lb : (Nat.cast n : ComputableℝSeq).lb = n := by 398 | rfl 399 | 400 | @[simp] 401 | theorem natCast_ub : (Nat.cast n : ComputableℝSeq).ub = n := by 402 | rfl 403 | 404 | @[simp] 405 | theorem val_natCast : (Nat.cast n : ComputableℝSeq).val = n := 406 | val_eq_mk_lb _ ▸ natCast_lb ▸ rfl 407 | 408 | @[simp] 409 | theorem intCast_lb : (Int.cast z : ComputableℝSeq).lb = z := by 410 | rfl 411 | 412 | @[simp] 413 | theorem intCast_ub : (Int.cast z : ComputableℝSeq).ub = z := by 414 | rfl 415 | 416 | @[simp] 417 | theorem val_intCast : (Int.cast z : ComputableℝSeq).val = z := 418 | val_eq_mk_lb _ ▸ intCast_lb ▸ rfl 419 | 420 | theorem ratCast_lb : (Rat.cast q : ComputableℝSeq).lb = CauSeq.const abs q := by 421 | rfl 422 | 423 | theorem ratCast_ub : (Rat.cast q : ComputableℝSeq).ub = CauSeq.const abs q := by 424 | rfl 425 | 426 | @[simp] 427 | theorem val_ratCast : (Rat.cast q : ComputableℝSeq).val = q := 428 | val_eq_mk_lb _ ▸ ratCast_lb ▸ rfl 429 | 430 | @[simp] 431 | theorem zero_lb : (0 : ComputableℝSeq).lb = 0 := by 432 | rfl 433 | 434 | @[simp] 435 | theorem zero_ub : (0 : ComputableℝSeq).ub = 0 := by 436 | rfl 437 | 438 | @[simp] 439 | theorem val_zero : (0 : ComputableℝSeq).val = 0 := 440 | val_eq_mk_lb _ ▸ Real.mk_zero 441 | 442 | @[simp] 443 | theorem one_lb : (1 : ComputableℝSeq).lb = 1 := by 444 | rfl 445 | 446 | @[simp] 447 | theorem one_ub : (1 : ComputableℝSeq).ub = 1 := by 448 | rfl 449 | 450 | @[simp] 451 | theorem val_one : (1 : ComputableℝSeq).val = 1 := 452 | val_eq_mk_lb _ ▸ Real.mk_one 453 | 454 | @[simp] 455 | theorem lb_add : (x + y).lb = x.lb + y.lb := 456 | rfl 457 | 458 | @[simp] 459 | theorem ub_add : (x + y).ub = x.ub + y.ub := 460 | rfl 461 | 462 | @[simp] 463 | theorem val_add : (x + y).val = x.val + y.val := by 464 | convert (mk_val_eq_val : (add x y).val = x.val + y.val) 465 | 466 | @[simp] 467 | theorem lb_neg : (-x).lb = -x.ub := 468 | rfl 469 | 470 | @[simp] 471 | theorem ub_neg : (-x).ub = -x.lb := by 472 | rfl 473 | 474 | @[simp] 475 | theorem val_neg : (-x).val = -x.val := by 476 | convert (mk_val_eq_val : (neg x).val = -x.val) 477 | 478 | @[simp] 479 | theorem lb_sub : (x - y).lb = x.lb - y.ub := by 480 | suffices (sub x y).lb = x.lb - y.ub by 481 | convert this 482 | rw [sub, add, neg] 483 | ext 484 | simp [mk, lb, ub, sub_eq_add_neg] 485 | 486 | @[simp] 487 | theorem ub_sub : (x - y).ub = x.ub - y.lb := by 488 | suffices (sub x y).ub = x.ub - y.lb by 489 | convert this 490 | rw [sub, add, neg] 491 | ext 492 | simp [mk, lb, ub, sub_eq_add_neg] 493 | 494 | @[simp] 495 | theorem val_sub : (x - y).val = x.val - y.val := by 496 | suffices (sub x y).val = x.val - y.val by 497 | convert this 498 | rw [sub, add, neg, mk_val_eq_val, mk_val_eq_val] 499 | rfl 500 | 501 | theorem lb_mul : (x * y).lb = ((x.lb * y.lb) ⊓ (x.ub * y.lb)) ⊓ ((x.lb * y.ub) ⊓ (x.ub * y.ub)) := by 502 | ext 503 | rw [← mul_lb, ← fst_mul'_eq_mul_lb] 504 | rfl 505 | 506 | theorem ub_mul : (x * y).ub = ((x.lb * y.lb) ⊔ (x.ub * y.lb)) ⊔ ((x.lb * y.ub) ⊔ (x.ub * y.ub)) := by 507 | ext 508 | rw [← mul_ub, ← snd_mul'_eq_mul_ub] 509 | rfl 510 | 511 | @[simp] 512 | theorem val_mul : (x * y).val = x.val * y.val := by 513 | suffices (mul x y).val = x.val * y.val by 514 | convert this 515 | rw [val_def] 516 | exact val_uniq' (mul_lb_is_lb x y) (mul_ub_is_ub x y) (lb_ub_mul_equiv x y) 517 | 518 | end simps 519 | 520 | section signs 521 | 522 | private noncomputable instance sign_aux_sound (x : ℝ) : 523 | Inhabited { s : SignType // s = SignType.sign x } := ⟨SignType.sign x, rfl⟩ 524 | 525 | /-- Compute the sign of x. Guaranteed to terminate if x is nonzero. If x is zero, it will 526 | terminate and return zero only if the lower and upper bounds become exactly zero at a finite n. 527 | Otherwise this becomes an infinite loop. For instance, `Real.pi - Real.pi` will never terminate. 528 | This ends up providing a `DecidableEq` and `DecidableLT` instance, but "in practice" this should 529 | only be used to prove nonequality, or check which of two inequal values is the larger -- not to 530 | prove equality. (The only equalities that will realistically end up being proven are the ones that 531 | could have been done entirely with rational numbers the whole way.) -/ 532 | partial def sign (x : ComputableℝSeq) : SignType := 533 | aux 0 where 534 | aux (n : ℕ) : { s : SignType // s = SignType.sign x.val } := 535 | let xun := x.ub n 536 | if h : xun < 0 then --upper bound is negative so x is negative 537 | ⟨SignType.neg, by rw [sign_neg]; rfl; rify at h; linarith [x.hub n] ⟩ 538 | else 539 | let xln := x.lb n 540 | if h₂ : xln > 0 then --lower bound is posiive so x is positive 541 | ⟨SignType.pos, by rw [sign_pos]; rfl; rify at h₂; linarith [x.hlb n] ⟩ 542 | else if h₃ : xln = 0 && xun = 0 then --x=0 exactly 543 | ⟨0, Eq.symm (by 544 | rw [sign_eq_zero_iff] 545 | simp only [Bool.and_eq_true, decide_eq_true_eq] at h₃ 546 | rify at h₃ 547 | linarith [x.hlb n, x.hub n] 548 | )⟩ 549 | else --not determined, proceed further in sequence 550 | aux (n+1) 551 | 552 | theorem sign_sound (x : ComputableℝSeq) : x.sign = SignType.sign x.val := 553 | (sign.aux x 0).property 554 | 555 | theorem sign_pos_iff (x : ComputableℝSeq) : x.sign = SignType.pos ↔ 0 < x.val := by 556 | rw [sign_sound, SignType.pos_eq_one, sign_eq_one_iff] 557 | 558 | theorem sign_neg_iff (x : ComputableℝSeq) : x.sign = SignType.neg ↔ x.val < 0 := by 559 | rw [sign_sound, SignType.neg_eq_neg_one, sign_eq_neg_one_iff] 560 | 561 | theorem sign_zero_iff (x : ComputableℝSeq) : x.sign = SignType.zero ↔ x.val = 0 := by 562 | rw [sign_sound, SignType.zero_eq_zero, sign_eq_zero_iff] 563 | 564 | /-- If x is nonzero, there is eventually a point in the Cauchy sequences where either the lower 565 | or upper bound prove this. This theorem states that this point exists. -/ 566 | noncomputable def sign_witness_term (x : ComputableℝSeq) (hnz : x.val ≠ 0) : 567 | { xq : ℕ × ℚ // (0:ℝ) < xq.2 ∧ xq.2 < abs x.val ∧ ∀ j ≥ xq.1, |(x.lb - x.ub) j| < xq.2} := by 568 | have hsx : abs x.val > 0 := by positivity 569 | have hq' : ∃(q:ℚ), (0:ℝ) < q ∧ q < abs x.val := exists_rat_btwn hsx 570 | obtain ⟨q, hq⟩ := Classical.indefiniteDescription _ hq' 571 | obtain ⟨hq₁, hq₂⟩ := hq 572 | obtain ⟨i, hi⟩ := Classical.indefiniteDescription _ (x.heq q (Rat.cast_pos.mp hq₁)) 573 | use (i, q) 574 | 575 | theorem sign_witness_term_prop (x : ComputableℝSeq) (n : ℕ) (hnz : x.val ≠ 0) 576 | (hub : ¬(x.ub).val n < 0) (hlb: ¬(x.lb).val n > 0) : 577 | n + Nat.succ 0 ≤ (x.sign_witness_term hnz).val.1 := by 578 | push_neg at hub hlb 579 | obtain ⟨⟨k, q⟩, ⟨h₁, h₂, h₃⟩⟩ := x.sign_witness_term hnz 580 | by_contra hn 581 | replace h₃ := h₃ n (by linarith) 582 | simp_rw [CauSeq.sub_apply] at h₃ 583 | rw [abs_ite_le] at h₂ h₃ 584 | have := x.hlb n 585 | have := x.hub n 586 | split_ifs at h₂ h₃ with h₄ h₅ 587 | all_goals 588 | rify at *; linarith (config := {splitNe := true}) 589 | 590 | /-- With the proof that x≠0, we can also eventually get a sign witness: a number n such that 591 | either 0 < x and 0 < lb n; or that x < 0 and ub n < 0. Marking it as irreducible because 592 | in theory all of the info needed is in the return Subtype. -/ 593 | irreducible_def sign_witness (x : ComputableℝSeq) (hnz : x.val ≠ 0) : 594 | { n // (0 < x.val ∧ 0 < x.lb n) ∨ (x.val < 0 ∧ x.ub n < 0)} := 595 | sign_witness_aux 0 hnz where 596 | sign_witness_aux (k : ℕ) (hnz : x.val ≠ 0) : { n // (0 < x.val ∧ 0 < x.lb n) ∨ (x.val < 0 ∧ x.ub n < 0)}:= 597 | if hub : x.ub k < 0 then 598 | ⟨k, Or.inr ⟨by rify at hub; linarith [x.hub k], hub⟩⟩ 599 | else if hlb : x.lb k > 0 then 600 | ⟨k, Or.inl ⟨by rify at hlb; linarith [x.hlb k], hlb⟩⟩ 601 | else 602 | sign_witness_aux (k+1) hnz 603 | termination_by 604 | (x.sign_witness_term hnz).val.fst - k 605 | decreasing_by 606 | · decreasing_with 607 | apply Nat.sub_add_lt_sub _ Nat.le.refl 608 | exact x.sign_witness_term_prop k hnz hub hlb 609 | 610 | /-- With the proof that x≠0, we get a total comparison function. -/ 611 | def is_pos {x : ComputableℝSeq} (hnz : x.val ≠ 0) : Bool := 612 | 0 < x.lb (x.sign_witness hnz) 613 | 614 | /-- Proof that `is_pos` correctly determines whether a nonzero computable number is positive. -/ 615 | theorem is_pos_iff (x : ComputableℝSeq) (hnz : x.val ≠ 0) : is_pos hnz ↔ 0 < x.val := by 616 | have hsw := (x.sign_witness hnz).property 617 | have hls := x.hlb (x.sign_witness hnz) 618 | have hus := x.hub (x.sign_witness hnz) 619 | constructor 620 | · intro h 621 | rw [is_pos, decide_eq_true_eq] at h 622 | cases hsw 623 | · tauto 624 | · rify at * 625 | linarith 626 | · intro h 627 | have := not_lt.mpr (le_of_lt h) 628 | rw [is_pos, decide_eq_true_eq] 629 | tauto 630 | 631 | theorem neg_of_not_pos {x : ComputableℝSeq} {hnz : x.val ≠ 0} (h : ¬is_pos hnz) : x.val < 0 := by 632 | rw [is_pos_iff] at h 633 | linarith (config := {splitNe := true}) 634 | 635 | /- Given computable sequences for a nonzero x, drop the leading terms of both sequences 636 | (lb and ub) until both are nonzero. This gives a new sequence that we can "safely" invert. 637 | -/ 638 | def dropTilSigned (x : ComputableℝSeq) (hnz : x.val ≠ 0) : ComputableℝSeq := 639 | let start := sign_witness x hnz 640 | mk (x := x.val) 641 | (lub := fun n ↦ x.lub (start+n)) 642 | (hcl := (x.lb.drop start).prop) 643 | (hcu := (x.ub.drop start).prop) 644 | (hlb := fun n ↦ x.hlb (start+n)) 645 | (hub := fun n ↦ x.hub (start+n)) 646 | (heq := Setoid.trans ( 647 | Setoid.trans (x.lb.drop_equiv_self start) x.heq) 648 | (Setoid.symm (x.ub.drop_equiv_self start))) 649 | 650 | @[simp] 651 | theorem val_dropTilSigned {x : ComputableℝSeq} (h : x.val ≠ 0) : (x.dropTilSigned h).val = x.val := by 652 | rw [val, val, Real.mk_eq] 653 | apply (lb x).drop_equiv_self 654 | 655 | theorem dropTilSigned_nz {x : ComputableℝSeq} (h : x.val ≠ 0) : (x.dropTilSigned h).val ≠ 0 := 656 | val_dropTilSigned h ▸ h 657 | 658 | theorem sign_dropTilSigned {x : ComputableℝSeq} (hnz : x.val ≠ 0) : 659 | (0 < x.val ∧ 0 < (x.dropTilSigned hnz).lb 0) ∨ (x.val < 0 ∧ (x.dropTilSigned hnz).ub 0 < 0) := by 660 | have := (x.sign_witness hnz).prop 661 | have := lt_trichotomy x.val 0 662 | tauto 663 | 664 | theorem dropTilSigned_pos {x : ComputableℝSeq} (h : x.val ≠ 0) : 665 | x.val > 0 ↔ (x.dropTilSigned h).lb 0 > 0 := 666 | ⟨fun h' ↦ (Or.resolve_right (sign_dropTilSigned h) 667 | (not_and.mpr fun a _ => ( not_lt_of_gt h') a)).2, 668 | fun h' ↦ calc val x = _ := (val_dropTilSigned h).symm 669 | _ ≥ _ := (x.dropTilSigned h).hlb 0 670 | _ > 0 := Rat.cast_pos.2 h'⟩ 671 | 672 | theorem dropTilSigned_neg {x : ComputableℝSeq} (h : x.val ≠ 0) : 673 | x.val < 0 ↔ (x.dropTilSigned h).ub 0 < 0 := 674 | ⟨fun h' ↦ (Or.resolve_left (sign_dropTilSigned h) 675 | (not_and.mpr fun a _ => ( not_lt_of_gt h') a)).2, 676 | fun h' ↦ calc val x = _ := (val_dropTilSigned h).symm 677 | _ ≤ _ := (x.dropTilSigned h).hub 0 678 | _ < 0 := Rat.cast_lt_zero.2 h'⟩ 679 | 680 | end signs 681 | 682 | section safe_inv 683 | 684 | theorem neg_LimZero_lb_of_val {x : ComputableℝSeq} (hnz : x.val ≠ 0) : ¬x.lb.LimZero := by 685 | rw [← CauSeq.Completion.mk_eq_zero] 686 | rw [val_eq_mk_lb, ←Real.mk_zero, ne_eq, Real.ofCauchy.injEq] at hnz 687 | exact hnz 688 | 689 | theorem neg_LimZero_ub_of_val {x : ComputableℝSeq} (hnz : x.val ≠ 0) : ¬x.ub.LimZero := by 690 | rw [← CauSeq.Completion.mk_eq_zero] 691 | rw [val_eq_mk_ub, ←Real.mk_zero, ne_eq, Real.ofCauchy.injEq] at hnz 692 | exact hnz 693 | 694 | theorem pos_ub_of_val {x : ComputableℝSeq} (hp : x.val > 0) : x.ub.Pos := 695 | Real.mk_pos.1 (val_eq_mk_ub _ ▸ hp) 696 | 697 | theorem pos_neg_ub_of_val {x : ComputableℝSeq} (hn : x.val < 0) : (-x.ub).Pos := 698 | Real.mk_pos.1 (lb_neg _ ▸ val_eq_mk_lb _ ▸ val_neg _ ▸ Left.neg_pos_iff.mpr hn) 699 | 700 | theorem pos_lb_of_val {x : ComputableℝSeq} (hp : x.val > 0) : x.lb.Pos := 701 | Real.mk_pos.1 (val_eq_mk_lb _ ▸ hp) 702 | 703 | theorem pos_neg_lb_of_val {x : ComputableℝSeq} (hn : x.val < 0) : (-x.lb).Pos := 704 | Real.mk_pos.1 (ub_neg _ ▸ val_eq_mk_ub _ ▸ val_neg _ ▸ Left.neg_pos_iff.mpr hn) 705 | 706 | /-- The sequence of lower bounds of 1/x. Only functions "correctly" to give lower bounds if we 707 | assume that hx is already `hx.dropTilSigned` (as proven in `lb_inv_correct`) -- but those 708 | assumptions aren't need for proving that it's Cauchy. -/ 709 | def lb_inv (x : ComputableℝSeq) (hnz : x.val ≠ 0) : CauSeq ℚ abs := 710 | if hp : is_pos hnz then --if x is positive, then reciprocals of ub's are always good lb's. 711 | x.ub.inv (neg_LimZero_ub_of_val hnz) 712 | else --x is negative, so positive values for ub don't give us any good lb's. 713 | let ub0 := x.ub 0 --save this first value, it acts as fallback if we get a bad ub 714 | ⟨fun n ↦ 715 | let ub := x.ub n 716 | if ub ≥ 0 then 717 | ub0⁻¹ --sign is indeterminate, fall back to the starting values 718 | else 719 | ub⁻¹, fun _ hε ↦ 720 | have hxv : x.val < 0 := by rw [is_pos_iff] at hp; linarith (config:={splitNe:=true}) 721 | let ⟨_, q0, Hq⟩ := pos_neg_ub_of_val hxv 722 | let ⟨_, K0, HK⟩ := CauSeq.abv_pos_of_not_limZero (neg_LimZero_ub_of_val hnz) 723 | let ⟨_, δ0, Hδ⟩ := rat_inv_continuous_lemma abs hε K0 724 | let ⟨i, H⟩ := exists_forall_ge_and (exists_forall_ge_and HK (x.ub.cauchy₃ δ0)) Hq 725 | let ⟨⟨iK, H'⟩, _⟩ := H _ le_rfl 726 | ⟨i, fun j hj ↦ 727 | have h₁ := CauSeq.neg_apply x.ub _ ▸ H _ le_rfl 728 | have h₁ := CauSeq.neg_apply x.ub _ ▸ H _ hj 729 | by 730 | simp only [(by linarith : ¬x.ub i ≥ 0),(by linarith : ¬x.ub j ≥ 0), ite_false] 731 | exact Hδ (H _ hj).1.1 iK (H' _ hj)⟩⟩ 732 | 733 | /-- Analgoous to `lb_inv` for providing upper bounds on 1/x. -/ 734 | def ub_inv (x : ComputableℝSeq) (hnz : x.val ≠ 0) : CauSeq ℚ abs := 735 | if hp : ¬is_pos hnz then --if x is positive, then reciprocals of ub's are always good lb's. 736 | x.lb.inv (neg_LimZero_lb_of_val hnz) 737 | else --x is negative, so positive values for ub don't give us any good lb's. 738 | let lb0 := x.lb 0 --save this first value, it acts as fallback if we get a bad ub 739 | ⟨fun n ↦ 740 | let lb := x.lb n 741 | if lb ≤ 0 then 742 | lb0⁻¹ --sign is indeterminate, fall back to the starting values 743 | else 744 | lb⁻¹, fun _ hε ↦ 745 | have hxv : x.val > 0 := by rw [is_pos_iff, not_not] at hp; linarith (config:={splitNe:=true}) 746 | let ⟨_, q0, Hq⟩ := pos_lb_of_val hxv 747 | let ⟨_, K0, HK⟩ := CauSeq.abv_pos_of_not_limZero (neg_LimZero_lb_of_val hnz) 748 | let ⟨_, δ0, Hδ⟩ := rat_inv_continuous_lemma abs hε K0 749 | let ⟨i, H⟩ := exists_forall_ge_and (exists_forall_ge_and HK (x.lb.cauchy₃ δ0)) Hq 750 | let ⟨⟨iK, H'⟩, _⟩ := H _ le_rfl 751 | ⟨i, fun j hj ↦ 752 | have h₁ := H _ le_rfl 753 | have h₁ := H _ hj 754 | by 755 | simp only [(by linarith : ¬x.lb i ≤ 0),(by linarith : ¬x.lb j ≤ 0), ite_false] 756 | exact Hδ (H j hj).1.1 iK (H' j hj) 757 | ⟩⟩ 758 | 759 | /-- When applied to a `dropTilSigned`, `lb_inv` is a correct lower bound on x⁻¹. -/ 760 | theorem lb_inv_correct {x : ComputableℝSeq} (hnz : x.val ≠ 0) : ∀n, 761 | (x.dropTilSigned hnz).lb_inv (dropTilSigned_nz hnz) n ≤ x.val⁻¹ := by 762 | intro n 763 | rw [lb_inv] 764 | split_ifs with hp 765 | · simp only [CauSeq.inv_apply, Rat.cast_inv] 766 | rw [is_pos_iff, val_dropTilSigned] at hp 767 | apply inv_anti₀ hp 768 | apply hub 769 | · have hv : val x < 0 := by rw [is_pos_iff, val_dropTilSigned] at hp; linarith (config:={splitNe:=true}) 770 | dsimp 771 | split_ifs with h 772 | <;> simp only [Rat.cast_inv] 773 | <;> apply (inv_le_inv_of_neg ?_ hv).2 (hub x _) 774 | · exact_mod_cast (dropTilSigned_neg hnz).1 hv 775 | · exact_mod_cast not_le.1 h 776 | 777 | /-- When applied to a `dropTilSigned`, `ub_inv` is a correct upper bound on x⁻¹. -/ 778 | theorem ub_inv_correct {x : ComputableℝSeq} (hnz : x.val ≠ 0) : ∀n, 779 | (x.dropTilSigned hnz).ub_inv (dropTilSigned_nz hnz) n ≥ x.val⁻¹ := by 780 | intro n 781 | rw [ub_inv] 782 | split_ifs with hp 783 | · have hv : val x > 0 := by rw [is_pos_iff, val_dropTilSigned] at hp; linarith (config:={splitNe:=true}) 784 | dsimp 785 | split_ifs with h 786 | <;> simp only [Rat.cast_inv] 787 | <;> apply inv_anti₀ ?_ ((val_dropTilSigned hnz) ▸ hlb _ _) 788 | · exact_mod_cast (dropTilSigned_pos hnz).1 hv 789 | · exact_mod_cast not_le.1 h 790 | · simp only [CauSeq.inv_apply, Rat.cast_inv] 791 | replace hp := val_dropTilSigned _ ▸ neg_of_not_pos hp 792 | rw [ge_iff_le, inv_le_inv_of_neg] 793 | · exact ((val_dropTilSigned hnz) ▸ hlb _ _) 794 | · exact hp 795 | · calc _ ≤ _ := ((val_dropTilSigned hnz) ▸ hlb _ _) 796 | _ < _ := hp 797 | 798 | /-- `x.lb_inv` converges to `(x.val)⁻¹`. -/ 799 | theorem lb_inv_converges {x : ComputableℝSeq} (hnz : x.val ≠ 0) : 800 | Real.mk (x.lb_inv hnz) = x.val⁻¹ := by 801 | apply Real.ext_cauchy 802 | rw [Real.cauchy_inv, Real.cauchy, Real.cauchy, Real.mk, val_eq_mk_ub, Real.mk, 803 | CauSeq.Completion.inv_mk (neg_LimZero_ub_of_val hnz), CauSeq.Completion.mk_eq, lb_inv] 804 | split_ifs with h 805 | · rfl 806 | · exact fun _ hε ↦ 807 | have hxv : x.val < 0 := by 808 | rw [is_pos_iff] at h 809 | linarith (config := {splitNe := true}) 810 | let ⟨q, q0, ⟨i, H⟩⟩ := pos_neg_ub_of_val hxv 811 | ⟨i, fun j hj ↦ 812 | have : ¬x.ub j ≥ 0 := by linarith [CauSeq.neg_apply x.ub _ ▸ H _ hj] 813 | by simp [this, hε]⟩ 814 | 815 | /-- When applied to a `dropTilSigned`, `lb_inv` is converges to x⁻¹. -/ 816 | theorem lb_inv_signed_converges {x : ComputableℝSeq} (hnz : x.val ≠ 0) : 817 | Real.mk ((x.dropTilSigned hnz).lb_inv (dropTilSigned_nz hnz)) = x.val⁻¹ := by 818 | simp [lb_inv_converges (dropTilSigned_nz hnz)] 819 | 820 | /-- `x.ub_inv` converges to `(x.val)⁻¹`. -/ 821 | theorem ub_inv_converges {x : ComputableℝSeq} (hnz : x.val ≠ 0) : 822 | Real.mk (x.ub_inv hnz) = x.val⁻¹ := by 823 | apply Real.ext_cauchy 824 | rw [Real.cauchy_inv, Real.cauchy, Real.cauchy, Real.mk, val_eq_mk_lb, Real.mk, 825 | CauSeq.Completion.inv_mk (neg_LimZero_lb_of_val hnz), CauSeq.Completion.mk_eq, ub_inv] 826 | split_ifs with h 827 | · exact fun _ hε ↦ 828 | have hxv : x.val > 0 := by 829 | rw [is_pos_iff] at h 830 | linarith (config := {splitNe := true}) 831 | let ⟨q, q0, ⟨i, H⟩⟩ := pos_lb_of_val hxv 832 | ⟨i, fun j hj ↦ 833 | have : ¬x.lb j ≤ 0 := by linarith [H _ hj] 834 | by simp [this, hε]⟩ 835 | · rfl 836 | 837 | /-- When applied to a `dropTilSigned`, `ub_inv` is converges to x⁻¹. 838 | TODO: version without hnz hypothesis. -/ 839 | theorem ub_inv_signed_converges {x : ComputableℝSeq} (hnz : x.val ≠ 0) : 840 | Real.mk ((x.dropTilSigned hnz).ub_inv (dropTilSigned_nz hnz)) = x.val⁻¹ := by 841 | simp [ub_inv_converges (dropTilSigned_nz hnz)] 842 | 843 | /- An inverse is defined only on reals that we can prove are nonzero. If we can prove they are 844 | nonzero, then we can prove that at some point we learn the sign, and so can start giving actual 845 | upper and lower bounds. There is a separate `inv` that uses `sign` to construct the proof of 846 | nonzeroness by searching along the sequence (but isn't guaranteed to terminate). -/ 847 | def safe_inv (x : ComputableℝSeq) (hnz : x.val ≠ 0) : ComputableℝSeq := 848 | --TODO currently this passes the sequence to lb_inv and ub_inv separately, which means we evaluate 849 | --things twice (and this can lead to exponential slowdown for long series of inverses). This should 850 | --be bundled 851 | let signed := x.dropTilSigned hnz 852 | let hnz' := val_dropTilSigned hnz ▸ hnz 853 | mk (x := x.val⁻¹) 854 | (lub := fun n ↦ ⟨⟨(signed.lb_inv hnz') n, (signed.ub_inv hnz') n⟩, 855 | Rat.cast_le.mp ((lb_inv_correct hnz n).trans (ub_inv_correct hnz n))⟩) 856 | (hcl := (signed.lb_inv hnz').2) 857 | (hcu := (signed.ub_inv hnz').2) 858 | (hlb := lb_inv_correct hnz) 859 | (hub := ub_inv_correct hnz) 860 | (heq := Real.mk_eq.1 ((lb_inv_signed_converges hnz).trans (ub_inv_signed_converges hnz).symm)) 861 | 862 | @[simp] 863 | theorem val_safe_inv {x : ComputableℝSeq} (hnz : x.val ≠ 0) : (x.safe_inv hnz).val = x.val⁻¹ := by 864 | rw [safe_inv, mk_val_eq_val] 865 | 866 | theorem val_safe_inv_ne_zero {x : ComputableℝSeq} (hnz : x.val ≠ 0) : (x.safe_inv hnz).val ≠ 0 := by 867 | rw [val_safe_inv, ne_eq, inv_eq_zero] 868 | exact hnz 869 | 870 | /-- Subtype of sequences with nonzero values. These admit a (terminating) inverse function. -/ 871 | def nzSeq := {x : ComputableℝSeq // x.val ≠ 0} 872 | 873 | def inv_nz : nzSeq → nzSeq := 874 | fun x ↦ ⟨x.val.safe_inv x.prop, val_safe_inv_ne_zero _⟩ 875 | 876 | @[simp] 877 | theorem val_inv_nz (x : nzSeq) : (inv_nz x).val.val = x.val.val⁻¹ := 878 | val_safe_inv _ 879 | 880 | instance instNzInv : Inv nzSeq := 881 | ⟨inv_nz⟩ 882 | 883 | end safe_inv 884 | 885 | section inv 886 | 887 | /-- Inverse of a computable real. Will terminate if the argument is nonzero, or if it is zero and the 888 | upper and lower bounds become exactly zero at some point. See `ComputableℝSeq.sign`. If you want 889 | to only call this in a way guaranteed to terminate, use `ComputableℝSeq.safe_inv`. -/ 890 | def inv : ComputableℝSeq → ComputableℝSeq := 891 | fun x ↦ match h : x.sign with 892 | | SignType.pos => x.safe_inv (x.sign_pos_iff.1 h).ne' 893 | | SignType.neg => x.safe_inv (x.sign_neg_iff.1 h).ne 894 | | SignType.zero => 0 895 | 896 | instance instInv : Inv ComputableℝSeq := 897 | ⟨inv⟩ 898 | 899 | instance instDiv : Div ComputableℝSeq := 900 | ⟨fun x y ↦ x * y⁻¹⟩ 901 | 902 | theorem inv_def (x : ComputableℝSeq) : x⁻¹ = x.inv := 903 | rfl 904 | 905 | /-- The inverse is equal to the `safe_inv`. This is an actual equality of sequences, not just equivalence. -/ 906 | theorem inv_eq_safe_inv {x : ComputableℝSeq} (hnz : x.val ≠ 0) : x⁻¹ = x.safe_inv hnz := by 907 | rw [inv_def, inv] 908 | split 909 | next h => rfl 910 | next h => rfl 911 | next h => 912 | absurd h 913 | rw [sign_zero_iff] 914 | exact hnz 915 | 916 | @[simp] 917 | theorem val_inv (x : ComputableℝSeq) : x⁻¹.val = x.val⁻¹ := by 918 | by_cases h : x.val = 0 919 | · rw [h, inv_zero, inv_def, inv] 920 | split 921 | next h => 922 | let _ := (x.sign_pos_iff.1 h).ne' 923 | contradiction 924 | next h => 925 | let _ := (x.sign_neg_iff.1 h).ne 926 | contradiction 927 | next h => exact val_zero 928 | · rwa [inv_eq_safe_inv, val_safe_inv] 929 | 930 | @[simp] 931 | theorem val_div (x y : ComputableℝSeq) : (x / y).val = x.val / y.val := by 932 | change (x * y⁻¹).val = x.val * y.val⁻¹ 933 | simp 934 | 935 | end inv 936 | 937 | section semiring --proving that computable real *sequences* form a commutative semiring 938 | 939 | theorem add_comm (x y: ComputableℝSeq) : x + y = y + x := by 940 | ext <;> simp only [ub_add, lb_add] <;> ring_nf 941 | 942 | theorem mul_comm (x y : ComputableℝSeq) : x * y = y * x := by 943 | ext n 944 | <;> simp only [lb_mul, ub_mul, mul_lb, mul_ub] 945 | · repeat rw [_root_.mul_comm (lb x)] 946 | repeat rw [_root_.mul_comm (ub x)] 947 | dsimp 948 | rw [inf_assoc, inf_assoc] 949 | congr 1 950 | rw [← inf_assoc, ← inf_assoc] 951 | nth_rw 2 [inf_comm] 952 | · repeat rw [_root_.mul_comm (lb x)] 953 | repeat rw [_root_.mul_comm (ub x)] 954 | dsimp 955 | rw [sup_assoc, sup_assoc] 956 | congr 1 957 | rw [← sup_assoc, ← sup_assoc] 958 | nth_rw 2 [sup_comm] 959 | 960 | /-TODO(mul_assoc) 961 | This theorem is annoying. When it's done, several other theorems follow too. They're all tagged TODO(mul_assoc). 962 | 963 | theorem mul_assoc (x y z : ComputableℝSeq) : (x * y) * z = x * (y * z) := by 964 | ext n 965 | · simp only [lb_mul, ub_mul, mul_lb, mul_ub] 966 | sorry 967 | · sorry 968 | 969 | theorem left_distrib (x y z : ComputableℝSeq) : x * (y + z) = x * y + x * z := by 970 | ext n 971 | <;> simp only [lb_mul, ub_mul, mul_lb, mul_ub, lb_add, ub_add] 972 | · dsimp 973 | simp only [_root_.left_distrib, add_inf, inf_add, inf_assoc] 974 | -- congr 1 975 | repeat sorry 976 | · sorry 977 | 978 | theorem right_distrib (x y z : ComputableℝSeq) : (x + y) * z = x * z + y * z := by 979 | rw [mul_comm, left_distrib, mul_comm, mul_comm z y] 980 | 981 | -/ 982 | 983 | theorem neg_mul (x y : ComputableℝSeq) : -x * y = -(x * y) := by 984 | ext 985 | · rw [lb_neg, lb_mul, ub_mul] 986 | simp only [lb_neg, ub_neg, CauSeq.coe_inf, CauSeq.coe_mul, CauSeq.coe_neg, neg_mul, 987 | Pi.inf_apply, Pi.neg_apply, Pi.mul_apply, CauSeq.neg_apply, CauSeq.coe_sup, Pi.sup_apply, neg_sup] 988 | nth_rewrite 2 [inf_comm] 989 | nth_rewrite 3 [inf_comm] 990 | ring_nf 991 | · rw [ub_neg, lb_mul, ub_mul] 992 | simp only [lb_neg, ub_neg, CauSeq.coe_inf, CauSeq.coe_mul, CauSeq.coe_neg, neg_mul, 993 | Pi.inf_apply, Pi.neg_apply, Pi.mul_apply, CauSeq.neg_apply, CauSeq.coe_sup, Pi.sup_apply, neg_inf] 994 | nth_rewrite 2 [sup_comm] 995 | nth_rewrite 3 [sup_comm] 996 | ring_nf 997 | 998 | theorem mul_neg (x y : ComputableℝSeq) : x * -y = -(x * y) := by 999 | rw [mul_comm, neg_mul, mul_comm] 1000 | 1001 | theorem neg_eq_of_add (x y : ComputableℝSeq) (h : x + y = 0) : -x = y := by 1002 | have hlb : ∀(x y : ComputableℝSeq), x + y = 0 → x.lb = -y.ub := by 1003 | intro x y h 1004 | ext n 1005 | let ⟨h₁, h₂⟩ := ComputableℝSeq.ext_iff.mp h 1006 | specialize h₁ n 1007 | specialize h₂ n 1008 | simp only [lb_add, ub_add, CauSeq.add_apply, zero_lb, zero_ub, CauSeq.zero_apply, CauSeq.neg_apply] at h₁ h₂ ⊢ 1009 | have h₃ := x.lb_le_ub n 1010 | have h₄ := y.lb_le_ub n 1011 | linarith (config := {splitNe := true}) 1012 | ext 1013 | · rw [lb_neg, CauSeq.neg_apply, hlb y x (add_comm _ _ ▸ h), CauSeq.neg_apply] 1014 | · rw [ub_neg, CauSeq.neg_apply, hlb x y h, CauSeq.neg_apply, neg_neg] 1015 | 1016 | /-- Computable sequences have *most* of the properties of a field, including negation, subtraction, 1017 | multiplication, division, IntCast all working as one would expect, with commutativity/associativity, 1018 | involutive negation, and distributive properties ... except for a few crucial facts that a - a ≠ 0, 1019 | a * a⁻¹ ≠ 1, and (a⁻¹)⁻¹ ≠ a. This typeclass collects all these facts together. 1020 | 1021 | TODO could include mul_inv_rev, inv_eq_of_mul, intCast_ofNat, intCast_negSucc. -/ 1022 | /-TODO(mul_assoc) 1023 | class CompSeqClass (G : Type u) extends 1024 | CommSemiring G, DivInvMonoid G, HasDistribNeg G, SubtractionCommMonoid G, IntCast G, RatCast G 1025 | -/ 1026 | 1027 | class CompSeqClass (G : Type u) extends 1028 | AddCommMonoid G, CommMagma G, MulZeroOneClass G, Inv G, Div G, 1029 | HasDistribNeg G, SubtractionCommMonoid G, NatCast G, IntCast G, RatCast G 1030 | 1031 | instance instSeqCompSeqClass : CompSeqClass ComputableℝSeq := by 1032 | refine' { 1033 | natCast := fun n => n 1034 | intCast := fun z => z 1035 | ratCast := fun q => q 1036 | zero := 0 1037 | one := 1 1038 | mul := (· * ·) 1039 | add := (· + ·) 1040 | neg := (- ·) 1041 | sub := (· - ·) 1042 | inv := (·⁻¹) 1043 | div := (· / ·) 1044 | nsmul := nsmulRec 1045 | zsmul := zsmulRec 1046 | --inline several of the "harder" proofs that can't be done automatically 1047 | add_comm := add_comm 1048 | mul_comm := mul_comm 1049 | neg_mul := neg_mul 1050 | mul_neg := mul_neg 1051 | neg_eq_of_add := neg_eq_of_add 1052 | 1053 | -- mul_assoc := mul_assoc 1054 | -- npow := @npowRec _ ⟨(1 : ℕ)⟩ ⟨(· * ·)⟩ 1055 | -- left_distrib := left_distrib 1056 | -- right_distrib := right_distrib 1057 | -- zpow := zpowRec 1058 | , .. } 1059 | all_goals 1060 | intros 1061 | first 1062 | | rfl 1063 | | ext 1064 | all_goals 1065 | try simp only [natCast_ub, natCast_lb, Nat.cast_add, Nat.cast_one, CauSeq.add_apply, CauSeq.one_apply, 1066 | CauSeq.zero_apply, CauSeq.neg_apply, lb_add, ub_add, one_ub, one_lb, zero_ub, zero_lb, ub_neg, 1067 | lb_neg, neg_add_rev, neg_neg, zero_add, add_zero] 1068 | try ring_nf 1069 | try rfl 1070 | try { 1071 | rename_i a n 1072 | simp only [lb_mul, ub_mul, zero_lb, zero_ub, mul_zero, zero_mul, one_lb, one_ub, mul_one, one_mul, 1073 | CauSeq.inf_idem, CauSeq.sup_idem, CauSeq.zero_apply, CauSeq.coe_inf, CauSeq.coe_sup, 1074 | Pi.sup_apply, Pi.inf_apply, sup_eq_right, inf_eq_left, lb_le_ub a n] 1075 | } 1076 | 1077 | /-TODO(mul_assoc) 1078 | @[simp] 1079 | theorem val_natpow (x : ComputableℝSeq) (n : ℕ): (x ^ n).val = x.val ^ n := by 1080 | induction n 1081 | · rw [pow_zero, val_one, pow_zero] 1082 | · rename_i ih 1083 | rw [pow_succ, pow_succ, val_mul, ih] 1084 | -/ 1085 | 1086 | end semiring 1087 | 1088 | 1089 | /-- The equivalence relation on ComputableℝSeq's given by converging to the same real value. -/ 1090 | instance equiv : Setoid (ComputableℝSeq) := 1091 | ⟨fun f g => f.val = g.val, 1092 | ⟨fun _ => rfl, Eq.symm, Eq.trans⟩⟩ 1093 | 1094 | theorem equiv_iff {x y : ComputableℝSeq} : x ≈ y ↔ x.val = y.val := 1095 | ⟨id, id⟩ 1096 | 1097 | end ComputableℝSeq 1098 | -------------------------------------------------------------------------------- /ComputableReal/ComputableReal.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.ComputableRSeq 2 | 3 | /-- Computable reals, defined as the quotient of ComputableℝSeq sequences -- sequences with 4 | Cauchy sequences of lower and upper bounds that converge to the same value -- by the equivalence 5 | relation of having the same converged value. This is similar to how reals are quotients of Cauchy 6 | sequence (without any guarantees on lower/upper bounds). -/ 7 | def Computableℝ := 8 | @Quotient ComputableℝSeq ComputableℝSeq.equiv 9 | 10 | namespace Computableℝ 11 | 12 | def mk : ComputableℝSeq → Computableℝ := 13 | Quotient.mk ComputableℝSeq.equiv 14 | 15 | def val : Computableℝ → ℝ := Quotient.lift ComputableℝSeq.val (fun _ _ h ↦ h) 16 | 17 | @[simp] 18 | theorem val_mk_eq_val : (mk x).val = x.val := 19 | rfl 20 | 21 | @[simp] 22 | theorem val_quot_eq_val (x : ComputableℝSeq) : val (⟦x⟧ : Computableℝ) = x.val := 23 | rfl 24 | 25 | @[simp] 26 | theorem eq_iff_seq_val (x y : ComputableℝSeq) : (mk x).val = (mk y).val ↔ mk x = mk y:= 27 | ⟨fun h ↦ Quotient.eq.2 h, Quotient.eq.1⟩ 28 | 29 | @[simp] 30 | theorem eq_iff_eq_val (x y : Computableℝ) : x.val = y.val ↔ x = y := 31 | ⟨by 32 | let ⟨x',hx'⟩ := Quotient.exists_rep x 33 | let ⟨y',hy'⟩ := Quotient.exists_rep y 34 | subst hx' 35 | subst hy' 36 | exact (eq_iff_seq_val x' y').1, 37 | congrArg val⟩ 38 | 39 | /-- Alternate version of mapℝ that doesn't directly refer to f₂, so it stays 40 | computable even if f₂ isn't. -/ 41 | def mapℝ' (f : ComputableℝSeq → ComputableℝSeq) (h : ∃ f₂ : ℝ → ℝ, ∀x, (f x).val = f₂ x.val) : 42 | Computableℝ → Computableℝ := 43 | Quotient.map f (fun a b h₂ ↦ h.elim fun _ h ↦ (h₂ ▸ h a).trans (h b).symm) 44 | 45 | /-- Given a unary function on sequences that clearly matches function on reals, lift it. -/ 46 | def mapℝ (f : ComputableℝSeq → ComputableℝSeq) {f₂ : ℝ → ℝ} (h : ∀x, (f x).val = f₂ x.val) : 47 | Computableℝ → Computableℝ := 48 | mapℝ' f ⟨f₂, h⟩ 49 | 50 | theorem mapℝ'_eq_mapℝ : mapℝ' f h = mapℝ f h₂ := by 51 | rfl 52 | 53 | /-- Alternate version of map₂ℝ that doesn't directly refer to f₂, so it stays 54 | computable even if f₂ isn't. -/ 55 | def map₂ℝ' (f : ComputableℝSeq → ComputableℝSeq → ComputableℝSeq) (h : ∃ f₂ : ℝ → ℝ → ℝ, ∀x y, 56 | (f x y).val = f₂ x.val y.val) : Computableℝ → Computableℝ → Computableℝ := 57 | Quotient.map₂ f (fun a b h₂ y z h₃ ↦ h.elim fun _ h ↦ (h₂ ▸ h₃ ▸ h a y).trans (h b z).symm) 58 | 59 | /-- Given a binary function that clearly mimics a standard real function, lift that. -/ 60 | def map₂ℝ (f : ComputableℝSeq → ComputableℝSeq → ComputableℝSeq) {f₂ : ℝ → ℝ → ℝ} 61 | (h : ∀x y, (f x y).val = f₂ x.val y.val) : 62 | Computableℝ → Computableℝ → Computableℝ := 63 | map₂ℝ' f ⟨f₂, h⟩ 64 | 65 | theorem map₂ℝ'_eq_map₂ℝ : map₂ℝ' f h = map₂ℝ f h₂ := by 66 | ext 67 | rw [map₂ℝ] 68 | 69 | @[simp] 70 | theorem val_mapℝ_eq_val : (@mapℝ f f₂ h x).val = f₂ x.val := by 71 | let ⟨x',hx'⟩ := Quotient.exists_rep x 72 | subst hx' 73 | rw [mapℝ, mapℝ', Quotient.map_mk] 74 | apply h 75 | 76 | @[simp] 77 | theorem val_map₂ℝ_eq_val : (@map₂ℝ f f₂ h x y).val = f₂ x.val y.val := by 78 | let ⟨x',hx'⟩ := Quotient.exists_rep x 79 | let ⟨y',hy'⟩ := Quotient.exists_rep y 80 | subst hx' 81 | subst hy' 82 | rw [map₂ℝ, map₂ℝ', Quotient.map₂_mk] 83 | apply h 84 | 85 | instance instComputableAdd : Add Computableℝ := 86 | ⟨map₂ℝ (· + ·) ComputableℝSeq.val_add⟩ 87 | 88 | instance instComputableMul : Mul Computableℝ := 89 | ⟨map₂ℝ (· * ·) ComputableℝSeq.val_mul⟩ 90 | 91 | instance instComputableNeg : Neg Computableℝ := 92 | ⟨mapℝ (- ·) ComputableℝSeq.val_neg⟩ 93 | 94 | instance instComputableSub : Sub Computableℝ := 95 | ⟨map₂ℝ (· - ·) ComputableℝSeq.val_sub⟩ 96 | 97 | instance instComputableZero : Zero Computableℝ := 98 | ⟨mk 0⟩ 99 | 100 | instance instComputableOne : One Computableℝ := 101 | ⟨mk 1⟩ 102 | 103 | variable (x y : Computableℝ) 104 | 105 | @[simp] 106 | theorem val_add : (x + y).val = x.val + y.val := 107 | val_map₂ℝ_eq_val 108 | 109 | @[simp] 110 | theorem val_mul : (x * y).val = x.val * y.val := 111 | val_map₂ℝ_eq_val 112 | 113 | @[simp] 114 | theorem val_neg : (-x).val = -(x.val) := 115 | val_mapℝ_eq_val 116 | 117 | @[simp] 118 | theorem val_sub : (x - y).val = x.val - y.val := 119 | val_map₂ℝ_eq_val 120 | 121 | @[simp] 122 | theorem val_zero : (0 : Computableℝ).val = 0 := by 123 | rw [Zero.toOfNat0, instComputableZero] 124 | simp only [val_mk_eq_val, ComputableℝSeq.val_zero] 125 | 126 | @[simp] 127 | theorem val_one : (1 : Computableℝ).val = 1 := by 128 | rw [One.toOfNat1, instComputableOne] 129 | simp only [val_mk_eq_val, ComputableℝSeq.val_one] 130 | 131 | theorem add_mk (x y : ComputableℝSeq) : mk x + mk y = mk (x + y) := 132 | rfl 133 | 134 | theorem mul_mk (x y : ComputableℝSeq) : mk x * mk y = mk (x * y) := 135 | rfl 136 | 137 | theorem sub_mk (x y : ComputableℝSeq) : mk x - mk y = mk (x - y) := 138 | rfl 139 | 140 | theorem neg_mk (x : ComputableℝSeq) : -mk x = mk (-x) := 141 | rfl 142 | 143 | instance instCommRing : CommRing Computableℝ := by 144 | refine' { natCast := fun n => mk n 145 | intCast := fun z => mk z 146 | zero := 0 147 | one := 1 148 | mul := (· * ·) 149 | add := (· + ·) 150 | neg := (- ·) 151 | sub := (· - ·) 152 | npow := npowRec --todo faster instances 153 | nsmul := nsmulRec 154 | zsmul := zsmulRec 155 | .. } 156 | all_goals 157 | intros 158 | first 159 | | rfl 160 | | rw [← eq_iff_eq_val] 161 | simp 162 | try ring_nf! 163 | 164 | @[simp] 165 | theorem val_natpow (x : Computableℝ) (n : ℕ): (x ^ n).val = x.val ^ n := by 166 | induction n 167 | · rw [pow_zero, val_one, pow_zero] 168 | · rename_i ih 169 | rw [pow_succ, pow_succ, val_mul, ih] 170 | 171 | @[simp] 172 | theorem val_nsmul (x : Computableℝ) (n : ℕ) : (n • x).val = n • x.val := by 173 | induction n 174 | · simp 175 | · rename_i ih 176 | simpa using ih 177 | 178 | section safe_inv 179 | 180 | set_option linter.dupNamespace false in 181 | private def nz_quot_equiv := Equiv.subtypeQuotientEquivQuotientSubtype 182 | (fun x : ComputableℝSeq ↦ x.val ≠ 0) 183 | (fun x : Computableℝ ↦ x ≠ 0) 184 | (fun _ ↦ ⟨ 185 | fun h h₂ ↦ by 186 | rw [← eq_iff_eq_val, val_zero] at h₂ 187 | exact h h₂, 188 | fun (h : ¬_ = 0) h₂ ↦ by 189 | rw [← eq_iff_eq_val, val_zero] at h 190 | exact h h₂⟩) 191 | (fun _ _ ↦ Iff.rfl) 192 | 193 | /-- Auxiliary inverse definition that operates on the nonzero Computableℝ values. -/ 194 | def safe_inv' : { x : Computableℝ // x ≠ 0 } → { x : Computableℝ // x ≠ 0 } := 195 | fun v ↦ nz_quot_equiv.invFun <| Quotient.map _ fun x y h₁ ↦ by 196 | change (ComputableℝSeq.inv_nz x).val.val = (ComputableℝSeq.inv_nz y).val.val 197 | rw [ComputableℝSeq.val_inv_nz x, ComputableℝSeq.val_inv_nz y, h₁] 198 | (nz_quot_equiv.toFun v) 199 | 200 | /-- Inverse of a nonzero Computableℝ, safe (terminating) as long as x is nonzero. -/ 201 | irreducible_def safe_inv (hnz : x ≠ 0) : Computableℝ := safe_inv' ⟨x, hnz⟩ 202 | 203 | @[simp] 204 | theorem safe_inv_val (hnz : x ≠ 0) : (x.safe_inv hnz).val = x.val⁻¹ := by 205 | let ⟨x',hx'⟩ := Quotient.exists_rep x 206 | subst hx' 207 | have : (nz_quot_equiv { val := ⟦x'⟧, property := hnz : { x : Computableℝ // x ≠ 0 } }) = 208 | ⟦{ val := x', property := (by 209 | rw [show (0 : Computableℝ) = ⟦0⟧ by rfl] at hnz 210 | contrapose! hnz 211 | rwa [← Computableℝ.val_zero, ← val_quot_eq_val, eq_iff_eq_val] at hnz 212 | )}⟧ := by 213 | apply Equiv.subtypeQuotientEquivQuotientSubtype_mk 214 | rw [safe_inv, safe_inv', val, Equiv.toFun_as_coe, Equiv.invFun_as_coe, Quotient.lift_mk, this, 215 | Quotient.map_mk, nz_quot_equiv, Equiv.subtypeQuotientEquivQuotientSubtype_symm_mk, 216 | Quotient.lift_mk, ComputableℝSeq.val_inv_nz] 217 | 218 | end safe_inv 219 | 220 | section field 221 | 222 | instance instComputableInv : Inv Computableℝ := 223 | ⟨mapℝ' (·⁻¹) ⟨(·⁻¹), ComputableℝSeq.val_inv⟩⟩ 224 | 225 | @[simp] 226 | theorem inv_val : (x⁻¹).val = (x.val)⁻¹ := by 227 | change (mapℝ' _ _ _).val = (x.val)⁻¹ 228 | rw [mapℝ'_eq_mapℝ] 229 | exact val_mapℝ_eq_val (h := ComputableℝSeq.val_inv) 230 | 231 | example : True := ⟨⟩ 232 | 233 | instance instField : Field Computableℝ := { instCommRing with 234 | qsmul := _ 235 | nnqsmul := _ 236 | exists_pair_ne := ⟨0, 1, by 237 | rw [ne_eq, ← eq_iff_eq_val, val_zero, val_one] 238 | exact zero_ne_one⟩ 239 | mul_inv_cancel := by 240 | intro a ha 241 | rw [← eq_iff_eq_val, val_mul, inv_val, val_one] 242 | have : val a ≠ 0 := by rwa [← val_zero, ne_eq, eq_iff_eq_val] 243 | field_simp 244 | inv_zero := by rw [← eq_iff_eq_val]; simp 245 | } 246 | 247 | @[simp] 248 | theorem div_val : (x / y).val = x.val / y.val := by 249 | change (x * y⁻¹).val = _ 250 | rw [val_mul, inv_val] 251 | field_simp 252 | 253 | end field 254 | 255 | section ordered 256 | 257 | variable (x y : Computableℝ) 258 | 259 | def lt : Prop := by 260 | apply Quotient.lift (fun z ↦ z.sign = SignType.pos) ?_ (y - x) 261 | intro a b h 262 | dsimp 263 | rw [ComputableℝSeq.sign_sound, ComputableℝSeq.sign_sound, h] 264 | 265 | instance instLT : LT Computableℝ := 266 | ⟨lt⟩ 267 | 268 | def le : Prop := by 269 | apply Quotient.lift (fun z ↦ SignType.zero ≤ z.sign) ?_ (y - x) 270 | intro a b h 271 | dsimp 272 | rw [ComputableℝSeq.sign_sound, ComputableℝSeq.sign_sound, h] 273 | 274 | instance instLE : LE Computableℝ := 275 | ⟨le⟩ 276 | 277 | @[simp] 278 | theorem lt_iff_lt : x.val < y.val ↔ x < y := by 279 | change x.val < y.val ↔ lt x y 280 | let ⟨x',hx'⟩ := Quotient.exists_rep x 281 | let ⟨y',hy'⟩ := Quotient.exists_rep y 282 | subst hx' 283 | subst hy' 284 | rw [lt, ← mk, val_mk_eq_val, val_mk_eq_val, sub_mk, mk, Quotient.lift_mk] 285 | rw [ComputableℝSeq.sign_pos_iff, ComputableℝSeq.val_sub, sub_pos] 286 | 287 | @[simp] 288 | theorem le_iff_le : x.val ≤ y.val ↔ x ≤ y := by 289 | change x.val ≤ y.val ↔ le x y 290 | let ⟨x',hx'⟩ := Quotient.exists_rep x 291 | let ⟨y',hy'⟩ := Quotient.exists_rep y 292 | subst hx' 293 | subst hy' 294 | rw [le, ← mk, val_mk_eq_val, val_mk_eq_val, sub_mk, mk, Quotient.lift_mk] 295 | rw [ComputableℝSeq.sign_sound, SignType.zero_eq_zero, sign_nonneg_iff] 296 | rw [ComputableℝSeq.val_sub, sub_nonneg] 297 | 298 | instance instDecidableLE : DecidableRel (fun (x y : Computableℝ) ↦ x ≤ y) := 299 | fun a b ↦ by 300 | change Decidable (le a b) 301 | rw [le] 302 | infer_instance 303 | 304 | --TODO: add a faster `min` and `max` that don't require sign computation. 305 | instance instLinearOrderedField : LinearOrderedField Computableℝ := by 306 | refine' { instField, instLT, instLE with 307 | decidableLE := inferInstance 308 | le_refl := _ 309 | le_trans := _ 310 | le_antisymm := _ 311 | add_le_add_left := _ 312 | zero_le_one := _ 313 | mul_pos := _ 314 | le_total := _ 315 | lt_iff_le_not_le := _ 316 | } 317 | all_goals 318 | intros 319 | simp only [← le_iff_le, ← lt_iff_lt, ← eq_iff_eq_val, val_add, val_mul, val_zero, val_one] at * 320 | first 321 | | linarith (config := {splitNe := true}) 322 | | apply mul_pos ‹_› ‹_› 323 | | apply le_total 324 | | apply lt_iff_le_not_le 325 | 326 | end ordered 327 | 328 | end Computableℝ 329 | -------------------------------------------------------------------------------- /ComputableReal/IsComputable.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.ComputableReal 2 | import Mathlib.Analysis.RCLike.Basic 3 | import Mathlib.Tactic.Peel 4 | 5 | /- Type class stating that `x:ℝ` has a ComputableℝSeq, i.e. that x is a computable number. Like 6 | `Decidable`, it carries data with it - even though (classically) we could prove that ever proposition 7 | is decidable, and every real is computable. -/ 8 | class IsComputable (x : ℝ) : Type where 9 | seq : ComputableℝSeq 10 | prop : seq.val = x 11 | 12 | namespace IsComputable 13 | 14 | /-- Turns one `IsComputable` into another one, given a proof that they're equal. This is directly 15 | analogous to `decidable_of_iff`, as a way to avoid `Eq.rec` on data-carrying instances. -/ 16 | def lift_eq {x y : ℝ} (h : x = y) : 17 | IsComputable x → IsComputable y := 18 | fun ⟨sx, hsx⟩ ↦ ⟨sx, h ▸ hsx⟩ 19 | 20 | def lift (fr : ℝ → ℝ) (fs : ComputableℝSeq → ComputableℝSeq) 21 | (h : ∀ a, (fs a).val = fr a.val) : 22 | IsComputable x → IsComputable (fr x) := 23 | fun ⟨sx, hsx⟩ ↦ ⟨fs sx, hsx ▸ h sx⟩ 24 | 25 | def lift₂ (fr : ℝ → ℝ → ℝ) (fs : ComputableℝSeq → ComputableℝSeq → ComputableℝSeq) 26 | (h : ∀a b, (fs a b).val = fr a.val b.val) : 27 | IsComputable x → IsComputable y → IsComputable (fr x y) := 28 | fun ⟨sx, hsx⟩ ⟨sy, hsy⟩ ↦ ⟨fs sx sy, hsx ▸ hsy ▸ h sx sy⟩ 29 | 30 | instance instComputableNat (n : ℕ) : IsComputable n := 31 | ⟨ComputableℝSeq.ofRat n, ComputableℝSeq.val_natCast⟩ 32 | 33 | instance instComputableInt (z : ℤ) : IsComputable z := 34 | ⟨ComputableℝSeq.ofRat z, ComputableℝSeq.val_intCast⟩ 35 | 36 | instance instComputableRat (q : ℚ) : IsComputable q := 37 | ⟨ComputableℝSeq.ofRat q, ComputableℝSeq.val_ratCast⟩ 38 | 39 | instance instComputableOfNat1 : IsComputable 40 | (@OfNat.ofNat.{0} Real 1 (@One.toOfNat1.{0} Real inferInstance)) := 41 | ⟨1, ComputableℝSeq.val_one⟩ 42 | 43 | instance instComputableOfNat0 : IsComputable 44 | (@OfNat.ofNat.{0} Real 0 (@Zero.toOfNat0.{0} Real inferInstance)) := 45 | ⟨0, ComputableℝSeq.val_zero⟩ 46 | 47 | instance instComputableOfNatAtLeastTwo : (n : ℕ) → [n.AtLeastTwo] → IsComputable ofNat(n) := 48 | fun _ _ ↦ ⟨ 49 | ⟨fun _ ↦ ⟨⟨_, _⟩, rfl.le⟩, 50 | IsCauSeq.const _, IsCauSeq.const _, 51 | fun _ ↦ by simpa using ⟨rfl.le, rfl.le⟩, 52 | by rfl⟩, 53 | ComputableℝSeq.val_eq_mk_lb _⟩ 54 | 55 | instance instComputableNeg (x : ℝ) [hx : IsComputable x] : IsComputable (-x) := 56 | lift _ (- ·) ComputableℝSeq.val_neg hx 57 | 58 | instance instComputableInv (x : ℝ) [hx : IsComputable x] : IsComputable (x⁻¹) := 59 | lift _ (·⁻¹) ComputableℝSeq.val_inv hx 60 | 61 | instance instComputableAdd [hx : IsComputable x] [hy : IsComputable y] : IsComputable (x + y) := 62 | lift₂ _ (· + ·) ComputableℝSeq.val_add hx hy 63 | 64 | instance instComputableSub [hx : IsComputable x] [hy : IsComputable y] : IsComputable (x - y) := 65 | lift₂ _ (· - ·) ComputableℝSeq.val_sub hx hy 66 | 67 | instance instComputableMul [hx : IsComputable x] [hy : IsComputable y] : IsComputable (x * y) := 68 | lift₂ _ (· * ·) ComputableℝSeq.val_mul hx hy 69 | 70 | instance instComputableDiv [hx : IsComputable x] [hy : IsComputable y] : IsComputable (x / y) := 71 | lift₂ _ (· / ·) ComputableℝSeq.val_div hx hy 72 | 73 | instance instComputableNatPow [hx : IsComputable x] (n : ℕ) : IsComputable (x ^ n) := by 74 | /-TODO(mul_assoc) 75 | Redo this to use native powering on the ComputableℝSeq. This avoids more costly 76 | (and inaccurate) interval multiplications. That will also turn it into exponentiation 77 | by squaring. 78 | -/ 79 | induction n 80 | · rw [pow_zero] 81 | infer_instance 82 | · rw [pow_succ] 83 | infer_instance 84 | 85 | instance instComputableZPow [hx : IsComputable x] (z : ℤ) : IsComputable (x ^ z) := by 86 | cases z 87 | · rw [Int.ofNat_eq_coe, zpow_natCast] 88 | infer_instance 89 | · simp only [zpow_negSucc] 90 | infer_instance 91 | 92 | instance instComputableNSMul [hx : IsComputable x] (n : ℕ) : IsComputable (n • x) := 93 | lift _ (n • ·) (by 94 | --TODO move to a ComputableℝSeq lemma 95 | intro a 96 | induction n 97 | · simp 98 | · rename_i ih 99 | simp [ih, succ_nsmul, add_mul] 100 | ) hx 101 | 102 | instance instComputableZSMul [hx : IsComputable x] (z : ℤ) : IsComputable (z • x) := by 103 | rw [zsmul_eq_mul] 104 | infer_instance 105 | 106 | instance instComputableQSMul [hx : IsComputable x] (q : ℚ) : IsComputable (q • x) := by 107 | change IsComputable (_ * _) 108 | infer_instance 109 | 110 | /-- When expressions involve that happen to be `IsComputable`, we can get a decidability 111 | instance by lifting them to a comparison on the `ComputableℝSeq`s, where comparison is 112 | computable. -/ 113 | instance instDecidableLE [hx : IsComputable x] [hy : IsComputable y] : Decidable (x ≤ y) := 114 | decidable_of_decidable_of_iff (p := Computableℝ.mk hx.seq ≤ Computableℝ.mk hy.seq) (by 115 | simp only [← Computableℝ.le_iff_le, Computableℝ.val_mk_eq_val, hx.prop, hy.prop] 116 | ) 117 | 118 | instance instDecidableEq [hx : IsComputable x] [hy : IsComputable y] : Decidable (x = y) := 119 | decidable_of_decidable_of_iff (p := (Computableℝ.mk hx.seq = Computableℝ.mk hy.seq)) (by 120 | simp only [← Computableℝ.eq_iff_eq_val, Computableℝ.val_mk_eq_val, hx.prop, hy.prop] 121 | ) 122 | 123 | instance instDecidableLT [hx : IsComputable x] [hy : IsComputable y] : Decidable (x < y) := 124 | decidable_of_decidable_of_iff (p := Computableℝ.mk hx.seq < Computableℝ.mk hy.seq) (by 125 | simp only [← Computableℝ.lt_iff_lt, Computableℝ.val_mk_eq_val, hx.prop, hy.prop] 126 | ) 127 | 128 | instance instDecidableLE_val (x y : ComputableℝSeq) : Decidable (x.val ≤ y.val) := 129 | @instDecidableLE x.val y.val ⟨x, rfl⟩ ⟨y,rfl⟩ 130 | 131 | instance instDecidableLT_val (x y : ComputableℝSeq) : Decidable (x.val < y.val) := 132 | @instDecidableLT x.val y.val ⟨x, rfl⟩ ⟨y,rfl⟩ 133 | 134 | example : ((3 : ℝ) + (5 : ℕ)) / 100 < (3 : ℚ) * (5 + (1 / 5)^2 - 1) ∧ 135 | (5:ℕ) = ((1:ℝ) + (2:ℚ)^2) := by 136 | native_decide 137 | 138 | end IsComputable 139 | 140 | 141 | /-- This is very similar to the statement 142 | `TendstoLocallyUniformly (fun n x ↦ (F n x : ℝ)) (fun (q : ℚ) ↦ f q) .atTop` 143 | but that only uses neighborhoods within the rationals, which is a strictly 144 | weaker condition. This uses neighborhoods in the ambient space, the reals. 145 | -/ 146 | def TendstoLocallyUniformlyWithout (F : ℕ → ℚ → ℚ) (f : ℝ → ℝ) : Prop := 147 | ∀ (ε : ℝ), 0 < ε → 148 | ∀ (x : ℝ), ∃ t ∈ nhds x, ∃ a, ∀ (b : ℕ), a ≤ b → ∀ (y : ℚ), ↑y ∈ t → 149 | |f y - ↑(F b y)| < ε 150 | 151 | theorem Real_mk_of_TendstoLocallyUniformly' (fImpl : ℕ → ℚ → ℚ) (f : ℝ → ℝ) 152 | (hfImpl : TendstoLocallyUniformlyWithout (fImpl) (f)) 153 | (hf : Continuous f) 154 | (x : CauSeq ℚ abs) 155 | : ∃ (h : IsCauSeq abs (fun n ↦ fImpl n (x n))), Real.mk ⟨_, h⟩ = f (.mk x) := by 156 | 157 | apply Real.of_near 158 | 159 | simp only [Metric.continuous_iff, gt_iff_lt, Real.dist_eq] at hf 160 | 161 | rcases x with ⟨x, hx⟩ 162 | intro ε hε 163 | 164 | obtain ⟨δ₁, hδ₁pos, hδ₁⟩ := hf (.mk ⟨x, hx⟩) _ (half_pos hε) 165 | obtain ⟨i₁, hi₁⟩ := cauchy_real_mk ⟨x, hx⟩ δ₁ hδ₁pos 166 | 167 | obtain ⟨i₂_nhd, hi₂_nhds, i₃, hi₃⟩ := hfImpl _ (half_pos hε) (.mk ⟨x, hx⟩) 168 | obtain ⟨nl, nu, ⟨hnl, hnu⟩, hnd_sub⟩ := mem_nhds_iff_exists_Ioo_subset.mp hi₂_nhds 169 | replace hnd_sub : ∀ (r : ℚ), nl < r ∧ r < nu → ↑r ∈ i₂_nhd := fun r a => hnd_sub a 170 | replace hi₃ : ∀ (b : ℕ), i₃ ≤ b → ∀ (y : ℚ), nl < y ∧ y < nu → |f ↑y - ↑(fImpl b y)| < (ε / 2) := by 171 | peel hi₃ 172 | exact fun h ↦ this (hnd_sub _ h) 173 | 174 | set ε_nhd := min (nu - (.mk ⟨x,hx⟩)) ((.mk ⟨x,hx⟩) - nl) with hε_nhd 175 | obtain ⟨i₄, hi₄⟩ := cauchy_real_mk ⟨x,hx⟩ (ε_nhd / 2) (by 176 | rw [hε_nhd, gt_iff_lt, ← min_div_div_right (zero_le_two), lt_inf_iff] 177 | constructor <;> linarith) 178 | 179 | have hεn₁ : ε_nhd ≤ _ := inf_le_left 180 | have hεn₂ : ε_nhd ≤ _ := inf_le_right 181 | 182 | set i := max i₁ (max i₃ i₄) with hi 183 | use i 184 | intro j hj 185 | simp only [hi, ge_iff_le, sup_le_iff] at hj 186 | 187 | specialize hδ₁ _ (hi₁ j (by linarith)) 188 | specialize hi₄ j (by linarith) 189 | specialize hi₃ j (by linarith) (x j) (by 190 | constructor 191 | · linarith [sub_le_of_abs_sub_le_left hi₄.le] 192 | · linarith [sub_le_of_abs_sub_le_right hi₄.le] 193 | ) 194 | 195 | calc |↑(fImpl j (x j)) - f (Real.mk ⟨x, hx⟩)| = 196 | |(↑(fImpl j (x j)) - f ↑(x j)) + (f ↑(x j) - f (Real.mk ⟨x, hx⟩))| := by congr; ring_nf 197 | _ ≤ |(↑(fImpl j (x j)) - f ↑(x j))| + |(f ↑(x j) - f (Real.mk ⟨x, hx⟩))| := abs_add _ _ 198 | _ < ε := by rw [abs_sub_comm]; linarith 199 | 200 | open scoped QInterval 201 | 202 | namespace ComputableℝSeq 203 | 204 | def of_TendstoLocallyUniformly_Continuous 205 | {f : ℝ → ℝ} (hf : Continuous f) 206 | (fImpl : ℕ → ℚInterval → ℚInterval) 207 | (fImpl_l : ℕ → ℚ → ℚ) 208 | (fImpl_u : ℕ → ℚ → ℚ) 209 | (hlb : ∀ (n : ℕ) (q : ℚInterval) (x : ℝ), x ∈ q → fImpl_l n q.fst ≤ f x) 210 | (hub : ∀ (n : ℕ) (q : ℚInterval) (x : ℝ), x ∈ q → f x ≤ fImpl_u n q.snd) 211 | (hImplDef : ∀ n q, fImpl n q = ⟨⟨fImpl_l n q.fst, fImpl_u n q.snd⟩, 212 | Rat.cast_le.mp ((hlb n q q.fst ⟨le_refl _, Rat.cast_le.mpr q.2⟩).trans 213 | (hub n q q.fst ⟨le_refl _, Rat.cast_le.mpr q.2⟩))⟩) 214 | (hTLU_l : TendstoLocallyUniformlyWithout fImpl_l f) 215 | (hTLU_u : TendstoLocallyUniformlyWithout fImpl_u f) 216 | (x : ComputableℝSeq) : ComputableℝSeq := 217 | mk 218 | (x := f x.val) 219 | (lub := fun n ↦ fImpl n (x.lub n)) 220 | (hcl := by 221 | obtain ⟨w, _⟩ := Real_mk_of_TendstoLocallyUniformly' fImpl_l f hTLU_l hf x.lb 222 | simp_rw [hImplDef] 223 | exact w 224 | ) 225 | (hcu := by 226 | obtain ⟨w, _⟩ := Real_mk_of_TendstoLocallyUniformly' fImpl_u f hTLU_u hf x.ub 227 | simp_rw [hImplDef] 228 | exact w 229 | ) 230 | (hlb := fun n ↦ by simp_rw [hImplDef]; exact hlb n (x.lub n) x.val ⟨x.hlb n, x.hub n⟩) 231 | (hub := fun n ↦ by simp_rw [hImplDef]; exact hub n (x.lub n) x.val ⟨x.hlb n, x.hub n⟩) 232 | (heq := by 233 | obtain ⟨_, h₁⟩ := Real_mk_of_TendstoLocallyUniformly' fImpl_l f hTLU_l hf x.lb 234 | obtain ⟨_, h₂⟩ := Real_mk_of_TendstoLocallyUniformly' fImpl_u f hTLU_u hf x.ub 235 | simp only [hImplDef, ← Real.mk_eq] 236 | rw [lb_eq_ub] at h₁ 237 | exact h₁.trans h₂.symm 238 | ) 239 | 240 | @[simp] 241 | theorem val_of_TendstoLocallyUniformly_Continuous (f) (hf : Continuous f) (fI fl fu h₁ h₂ h₃ h₄ h₅ a) 242 | : (of_TendstoLocallyUniformly_Continuous hf fI fl fu h₁ h₂ h₃ h₄ h₅ a).val = 243 | f a.val := 244 | ComputableℝSeq.mk_val_eq_val 245 | 246 | end ComputableℝSeq 247 | 248 | --mostly a helper for numerical investigation 249 | def Rat.toDecimal (q : ℚ) (prec : ℕ := 20) := 250 | let p : ℚ → String := fun q ↦ 251 | (q.floor.repr).append <| ".".append <| (10^prec * (q - q.floor)).floor.repr.leftpad prec '0' 252 | if q < 0 then "-".append (p (-q)) else if q = 0 then "0" else p q 253 | -------------------------------------------------------------------------------- /ComputableReal/IsComputableC.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.IsComputable 2 | import ComputableReal.SpecialFunctions.Sqrt 3 | import Mathlib.Analysis.InnerProductSpace.Basic 4 | 5 | /- Type class stating that `x:ℂ` has a ComputableℝSeq for its real and imaginary parts. 6 | Note that we can't define this as `IsComputable x.re`+`IsComputable x.im`, because then 7 | (if `x` is a noncomputable expression), this will be a noncomputable expression. -/ 8 | -- class IsComputableℂ (x : ℂ) : Type where 9 | -- re : ComputableℝSeq 10 | -- im : ComputableℝSeq 11 | -- -- prop_re : re.val = x.re 12 | -- -- prop_im : im.val = x.im] 13 | class IsComputableℂ (x : ℂ) : Type where 14 | re : IsComputable x.re 15 | im : IsComputable x.im 16 | 17 | namespace IsComputableℂ 18 | 19 | open ComplexConjugate 20 | 21 | /-- Turns one `IsComputableℂ` into another one, given a proof that they're equal. This is directly 22 | analogous to `decidable_of_iff`, as a way to avoid `Eq.rec` on data-carrying instances. -/ 23 | def lift_eq {x y : ℂ} (h : x = y) [hx : IsComputableℂ x] : IsComputableℂ y := 24 | ⟨hx.1.lift_eq (congrArg _ h), hx.2.lift_eq (congrArg _ h)⟩ 25 | 26 | --We'll need some version of this once we want nontrivial functions, like exp/sin. 27 | 28 | -- def lift (fr : ℝ → ℝ) (fs : ComputableℝSeq → ComputableℝSeq) 29 | -- (h : ∀ a, (fs a).val = fr a.val) : 30 | -- IsComputable x → IsComputable (fr x) := 31 | -- fun ⟨sx, hsx⟩ ↦ ⟨fs sx, hsx ▸ h sx⟩ 32 | 33 | -- def lift₂ (fr : ℝ → ℝ → ℝ) (fs : ComputableℝSeq → ComputableℝSeq → ComputableℝSeq) 34 | -- (h : ∀a b, (fs a b).val = fr a.val b.val) : 35 | -- IsComputable x → IsComputable y → IsComputable (fr x y) := 36 | -- fun ⟨sx, hsx⟩ ⟨sy, hsy⟩ ↦ ⟨fs sx sy, hsx ▸ hsy ▸ h sx sy⟩ 37 | 38 | variable (x y : ℂ) [hx : IsComputableℂ x] [hy : IsComputableℂ y] 39 | 40 | variable (r : ℝ) [hr : IsComputable r] 41 | 42 | instance instComputableRe : IsComputable x.re := 43 | hx.re 44 | 45 | instance instComputableIm : IsComputable x.im := 46 | hx.im 47 | 48 | instance instComputableRI [hi : IsComputable i] : IsComputableℂ ⟨r, i⟩ := 49 | ⟨hr, hi⟩ 50 | 51 | instance instComputableI : IsComputableℂ .I := 52 | ⟨inferInstanceAs (IsComputable 0), inferInstanceAs (IsComputable 1)⟩ 53 | 54 | instance instComputableNat (n : ℕ) : IsComputableℂ n := 55 | ⟨inferInstanceAs (IsComputable n), inferInstanceAs (IsComputable 0)⟩ 56 | 57 | instance instComputableInt (z : ℤ) : IsComputableℂ z := 58 | ⟨inferInstanceAs (IsComputable z), inferInstanceAs (IsComputable 0)⟩ 59 | 60 | --TODO drop noncomputable when bump mathlib 61 | noncomputable instance instComputableRat (q : ℚ) : IsComputableℂ q := 62 | ⟨inferInstanceAs (IsComputable q), inferInstanceAs (IsComputable 0)⟩ 63 | 64 | instance instComputableReal : IsComputableℂ r := 65 | ⟨hr, inferInstanceAs (IsComputable 0)⟩ 66 | 67 | instance instComputableOfNat0 : IsComputableℂ 68 | (@OfNat.ofNat.{0} ℂ 0 (@Zero.toOfNat0.{0} ℂ inferInstance)) := 69 | ⟨inferInstanceAs (IsComputable 0), inferInstanceAs (IsComputable 0)⟩ 70 | 71 | instance instComputableOfNat1 : IsComputableℂ 72 | (@OfNat.ofNat.{0} ℂ 1 (@One.toOfNat1.{0} ℂ inferInstance)) := 73 | ⟨inferInstanceAs (IsComputable 1), inferInstanceAs (IsComputable 0)⟩ 74 | 75 | instance instComputableOfNatAtLeastTwo (n : ℕ) [n.AtLeastTwo] : IsComputableℂ ofNat(n) := 76 | ⟨inferInstanceAs (IsComputable n), inferInstanceAs (IsComputable 0)⟩ 77 | 78 | instance instComputableNeg : IsComputableℂ (-x) := 79 | ⟨let _ := hx.1; inferInstanceAs (IsComputable (-x.re)), 80 | let _ := hx.2; inferInstanceAs (IsComputable (-x.im))⟩ 81 | 82 | instance instComputableAdd : IsComputableℂ (x + y) := 83 | ⟨let _ := hx.1; let _ := hy.1; inferInstanceAs (IsComputable (x.re + y.re)), 84 | let _ := hx.2; let _ := hy.2; inferInstanceAs (IsComputable (x.im + y.im))⟩ 85 | 86 | instance instComputableSub : IsComputableℂ (x - y) := 87 | ⟨let _ := hx.1; let _ := hy.1; inferInstanceAs (IsComputable (x.re - y.re)), 88 | let _ := hx.2; let _ := hy.2; inferInstanceAs (IsComputable (x.im - y.im))⟩ 89 | 90 | /- 91 | TODO: Multiplication uses each component, twice; this means that long product expressions 92 | of complex numbers experience an exponential slowdown. Since we don't do any kind of caching, 93 | the correct solution would be to make `IsComputableℂ` actually carry a `ℕ → QInterval × QInterval` 94 | for evaluation, together with proofs that each half is a `ComputableℝSeq`. (Or to define more 95 | new types, like a `ComputableℂSeq` that captures all this.) 96 | -/ 97 | @[inline] 98 | instance instComputableMul : IsComputableℂ (x * y) := 99 | let _ := hx.1; 100 | let _ := hy.1; 101 | let _ := hx.2; 102 | let _ := hy.2; 103 | ⟨inferInstanceAs (IsComputable (_ - _)), inferInstanceAs (IsComputable (_ + _))⟩ 104 | 105 | instance instComputableNatPow (n : ℕ) : IsComputableℂ (x ^ n) := by 106 | /- TODO do this by exponentation by squaring -/ 107 | induction n 108 | · rw [pow_zero] 109 | infer_instance 110 | · rw [pow_succ] 111 | infer_instance 112 | 113 | /- 114 | This could be just `inferInstanceAs (IsComputable (_ + _))`, because `Complex.normSq` 115 | is defeq to `x.re * x.re + x.im * x.im`. But doing that naively means we evaluate `x.re` 116 | and `x.im` each twice. If (when?) we get a faster implementation of natpow, this will 117 | be more efficient. 118 | -/ 119 | instance instComputableNormSq : IsComputable (Complex.normSq x) := 120 | .lift_eq (x := x.re ^ 2 + x.im ^ 2) (by rw [Complex.normSq, pow_two, pow_two]; rfl) 121 | inferInstance 122 | 123 | instance instComputableStar : IsComputableℂ (conj x) := 124 | ⟨hx.1, let _ := hx.2; inferInstanceAs (IsComputable (-x.im))⟩ 125 | 126 | instance instComputableInv : IsComputableℂ (x⁻¹) := 127 | ⟨let _ := hx.1; .lift_eq (Complex.inv_re x).symm inferInstance, 128 | let _ := hx.2; .lift_eq (Complex.inv_im x).symm inferInstance⟩ 129 | 130 | instance instComputableDiv : IsComputableℂ (x / y) := 131 | inferInstanceAs (IsComputableℂ (_ * _)) 132 | 133 | instance instComputableZPow (z : ℤ) : IsComputableℂ (x ^ z) := 134 | z.casesOn 135 | (fun a ↦ inferInstanceAs (IsComputableℂ (x ^ a))) 136 | (fun a ↦ inferInstanceAs (IsComputableℂ (x ^ a.succ)⁻¹)) 137 | 138 | instance instComputableNSMul (n : ℕ) : IsComputableℂ (n • x) := 139 | ⟨let _ := hx.1; .lift_eq (Complex.re_nsmul n x).symm inferInstance, 140 | let _ := hx.2; .lift_eq (Complex.im_nsmul n x).symm inferInstance⟩ 141 | 142 | --Need it to pick the right instance for the smul. Probably a better way to do this 143 | attribute [-instance] Complex.instNormedAddCommGroup in 144 | attribute [-instance] Complex.instNormedField in 145 | attribute [-instance] Complex.instDenselyNormedField in 146 | attribute [-instance] Complex.instRCLike in 147 | instance instComputableZSMul (z : ℤ) : IsComputableℂ (z • x) := 148 | ⟨let _ := hx.1; .lift_eq (Complex.re_zsmul z x).symm inferInstance, 149 | let _ := hx.2; .lift_eq (Complex.im_zsmul z x).symm inferInstance⟩ 150 | 151 | --TODO: Can't find a way to make this computable (it really wants to use Complex.instField) 152 | noncomputable instance instComputableQSMul (q : ℚ) : IsComputableℂ (q • x) := 153 | lift_eq (Rat.smul_def q x).symm 154 | --Alternative: 155 | -- ⟨let _ := hx.1; .lift_eq (Complex.re_qsmul q x).symm inferInstance, 156 | -- let _ := hx.2; .lift_eq (Complex.im_qsmul q x).symm inferInstance⟩ 157 | 158 | instance instComputableInner : IsComputable (inner x y) := 159 | inferInstanceAs (IsComputable (Complex.re (conj x * y))) 160 | 161 | instance instComputableNorm : IsComputable ‖x‖ := 162 | inferInstanceAs (IsComputable (√(Complex.normSq x))) 163 | 164 | instance instComputableNNNorm : IsComputable ‖x‖₊ := 165 | instComputableNorm x 166 | 167 | /- 168 | To prove that two complex numbers are not equal, we could use `ext`, but it's better to 169 | compute a series for Complex.normSq. Otherwise `Real.sqrt 2 + I` and `Real.sqrt 2 - I` 170 | will never be comparable, because comparing the real parts will never terminate; but 171 | comparing the the norm of their difference will eventually be lower bounded. 172 | -/ 173 | instance instDecidableEq : Decidable (x = y) := 174 | decidable_of_decidable_of_iff (p := Complex.normSq (x - y) = 0) 175 | (by rw [Complex.normSq_eq_zero, sub_eq_zero]) 176 | 177 | open ComplexOrder in 178 | instance instDecidableLE : Decidable (x ≤ y) := 179 | inferInstanceAs (Decidable (_ ∧ _)) 180 | 181 | open ComplexOrder in 182 | instance instDecidableLT : Decidable (x < y) := 183 | inferInstanceAs (Decidable (_ ∧ _)) 184 | 185 | example : (1 + Complex.I) * (1 - Complex.I : ℂ) = 2 := by 186 | native_decide 187 | 188 | example : ‖Complex.I‖ ≠ (1 / 2) := by 189 | native_decide 190 | -------------------------------------------------------------------------------- /ComputableReal/SpecialFunctions.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.SpecialFunctions.Basic 2 | import ComputableReal.SpecialFunctions.Sqrt 3 | import ComputableReal.SpecialFunctions.Pi 4 | import ComputableReal.SpecialFunctions.Exp 5 | -------------------------------------------------------------------------------- /ComputableReal/SpecialFunctions/Basic.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.IsComputable 2 | import Mathlib.Data.Real.Sign 3 | import Mathlib.Data.Real.ConjExponents 4 | 5 | namespace IsComputable 6 | 7 | instance instComputableOfScientific (m : ℕ) (b : Bool) (e : ℕ) : 8 | IsComputable (@OfScientific.ofScientific ℝ NNRatCast.toOfScientific m b e) := 9 | ⟨(OfScientific.ofScientific m b e : ℚ), 10 | ComputableℝSeq.val_ratCast.trans (Rat.cast_ofScientific m b e)⟩ 11 | 12 | --Note that if `(x y : ℝ)` comes after `Decidable`, the compiler won't drop them 13 | -- early enough in the compilation process, and this will fail as noncomputable. 14 | instance instComputableDite (p : Prop) (x : p → ℝ) (y : ¬p → ℝ) [Decidable p] 15 | [hx : ∀ p, IsComputable (x p)] [hy : ∀ p, IsComputable (y p)] 16 | : IsComputable (dite p x y) := 17 | if h : p then lift_eq (by simp [h]) (hx h) else lift_eq (by simp [h]) (hy h) 18 | 19 | @[inline] 20 | instance instComputableIte (p : Prop) (x y : ℝ) [Decidable p] 21 | [hx : IsComputable x] [hy : IsComputable y] : IsComputable (ite p x y) := 22 | instComputableDite p _ _ 23 | 24 | @[inline] 25 | instance instComputableSign (x : ℝ) [hx : IsComputable x] : IsComputable (x.sign) := 26 | ⟨hx.seq.sign, 27 | by 28 | rw [ComputableℝSeq.sign_sound, hx.prop] 29 | rcases lt_trichotomy x 0 with h|h|h 30 | · simp [Real.sign.eq_1, h] 31 | · simp [h] 32 | · simp [Real.sign.eq_1, h, not_lt_of_gt h]⟩ 33 | 34 | --TODO: This really should operate "directly" on the sequences, so that it doesn't 35 | --require a comparison first. For instance, `max (√2 - √2) 0 < 1` will never terminate 36 | --with this implementation. 37 | instance instComputableMax (x y : ℝ) [hx : IsComputable x] [hy : IsComputable y] : 38 | IsComputable (max x y) := 39 | lift_eq (x := ite (x ≤ y) ..) (by rw [max_def]; congr) inferInstance 40 | 41 | instance instComputableMin (x y : ℝ) [hx : IsComputable x] [hy : IsComputable y] : 42 | IsComputable (min x y) := 43 | lift_eq (x := ite (x ≤ y) ..) (by rw [min_def]; congr) inferInstance 44 | 45 | --This ends up calling the same sequence twice (which leads to an exponential 46 | --slowdown when many nested `abs` are present); would be good to write one that 47 | --directly takes the abs of each interval. Also, never returns a value for 48 | -- `abs (√2 - √2)`, for the same reasons as max. 49 | instance instComputableAbs (x : ℝ) [hx : IsComputable x] : IsComputable |x| := 50 | lift_eq (abs.eq_1 x).symm inferInstance 51 | 52 | instance instComputableConjExponent (x : ℝ) [hx : IsComputable x] : 53 | IsComputable x.conjExponent := 54 | lift_eq (Real.conjExponent.eq_1 x).symm inferInstance 55 | 56 | --Again, the order of arguments matters: if the argument `xs` is moved after 57 | -- `s`, then compilation breaks. 58 | instance instComputableFinsetSum {ι : Type*} (xs : ι → ℝ) (s : Finset ι) 59 | [hxs : ∀ i, IsComputable (xs i)] : IsComputable (Finset.sum s xs) := 60 | ⟨Finset.fold (· + ·) 0 (fun i ↦ (hxs i).seq) s, 61 | by 62 | rw [← Finset.sum_eq_fold] 63 | trans (∑ x ∈ s, (seq (xs x)).val) 64 | · --TODO: pull this out to be top-level theorems (that `.val` is an `AddMonoidHom`) 65 | exact map_sum (G := ComputableℝSeq →+ ℝ) 66 | ⟨⟨ComputableℝSeq.val, ComputableℝSeq.val_zero⟩, ComputableℝSeq.val_add⟩ 67 | (fun x ↦ seq (xs x)) s 68 | · congr! with i hi 69 | exact (hxs i).prop 70 | ⟩ 71 | 72 | --We would like to have this for `Finset.prod` as well, but the analogous proof of that 73 | --requires the fact that `ComputableℝSeq` multiplication is associative! Which we don't 74 | --have at the moment. 75 | 76 | end IsComputable 77 | 78 | example : 79 | (10 / 9 : ℝ) < (15 / 2 ⊔ 3) ∧ 80 | |(1 / 10) - (1 : ℝ)| < 1 ∧ 81 | (if _ : 1 < 2 then 5 / 2 else 6 : ℝ) < 7 ∧ 82 | (if [1,10].length = 2 then 5 else (10 / 7) : ℝ) / 7 < 1 ∧ 83 | (3.5 : ℝ) < 4 84 | := by 85 | native_decide 86 | 87 | example : (2 - 5 / 2 : ℝ).sign + 1 = 0 := by 88 | native_decide 89 | 90 | example : |∑ x ∈ Finset.range 500, 1 / (x : ℝ) - 6.7908234| < 0.0000001 := by 91 | native_decide 92 | 93 | /- 94 | There's some very nasty behavior where the compiler may or may not compile something 95 | computably, depending on the exact expression complexity. 96 | 97 | This works: 98 | #eval ((10 / 9 : ℝ) < (15 / 2 ⊔ 3) ∧ 99 | |(1 / 10) - (1 : ℝ)| < 1 ∧ 100 | (if _ : 1 < 2 then 5 / 2 else 6 : ℝ) < 7 ∧ 101 | (if [1,10].length = 2 then 5 else (10 / 7) : ℝ) / 7 < 1 ∧ 102 | (3.5 : ℝ) < 4) 103 | 104 | This works: 105 | #eval ((2 - 5 / 2 : ℝ).sign + 1 = 0) 106 | 107 | But this doesn't: 108 | #eval ((10 / 9 : ℝ) < (15 / 2 ⊔ 3) ∧ 109 | |(1 / 10) - (1 : ℝ)| < 1 ∧ 110 | (if _ : 1 < 2 then 5 / 2 else 6 : ℝ) < 7 ∧ 111 | (if [1,10].length = 2 then 5 else (10 / 7) : ℝ) / 7 < 1 ∧ 112 | (3.5 : ℝ) < 4 ∧ 113 | (2 - 5 / 2 : ℝ).sign + 1 = 0) 114 | 115 | -/ 116 | -------------------------------------------------------------------------------- /ComputableReal/SpecialFunctions/Exp.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.IsComputable 2 | import Mathlib.Data.Complex.Exponential 3 | import Mathlib.Analysis.SpecialFunctions.Exp 4 | import Mathlib.Analysis.SpecialFunctions.Pow.Real 5 | 6 | namespace ComputableℝSeq 7 | namespace Exp 8 | 9 | /- 10 | The Taylor series error formula, where fN is the nth-order approximation: 11 | f(x) - fN(x) = 1/(n+1)! * f^(n+1)(c) * x^(n+1) for some c in [0,x]. 12 | 13 | For exp: 14 | exp(x) - expN(x) = 1/(n+1)! * exp(c) * x^(n+1) 15 | 16 | |exp(x) - expN(x)| = |1/(n+1)! * exp(c) * x^(n+1)| 17 | <= |1/(n+1)! * exp(x) * x^(n+1)| 18 | 19 | |expN(x) - exp(x)| <= 1/(n+1)! * exp(x) * x^(n+1) 20 | 21 | expN(x) <= exp(x) + 1/(n+1)! * exp(x) * x^(n+1) 22 | 23 | expN(x) <= exp(x) (1 + x^(n+1) / (n+1)!) 24 | 25 | ∴ exp(x) >= expN(x) / (1 + x^(n+1) / (n+1)!) 26 | 27 | likewise, 28 | 29 | exp(x) <= expN(x) / (1 - x^(n+1) / (n+1)!) 30 | if (1 - x^(n+1) / (n+1)!) >= 0, otherwise 0 31 | -/ 32 | 33 | /-- A valid lower bound when 0 ≤ x. Forms a `CauSeq` that converges to Real.exp x from below. 34 | Functions by dividing `n` by a constant `k` so that it's in the range `[0,1]`, taking `n` terms 35 | of the Taylor expansion (which is an under-approximation), and then raising it to `k` again. -/ 36 | def exp_lb₀ (x : ℚ) (n : ℕ) : ℚ := 37 | let xc := ⌈x⌉.toNat 38 | let y : ℚ := x / xc 39 | -- (Finset.sum (Finset.range n) fun i => x ^ i / ↑(Nat.factorial i)) 40 | let ey : ℚ := (List.range n).foldr (fun (n : ℕ) v ↦ 1 + (y * v) / (n+1)) 1 41 | ey ^ xc 42 | 43 | /-- A valid upper bound when 0 ≤ x. CauSeq that converges to Real.exp x from above. -/ 44 | def exp_ub₀ (x : ℚ) (n : ℕ) : ℚ := 45 | let xc := ⌈x⌉.toNat 46 | let y : ℚ := x / xc 47 | -- (Finset.sum (Finset.range n) fun i => x ^ i / ↑(Nat.factorial i)) 48 | let ey : ℚ := (List.range n).foldr (fun (n : ℕ) v ↦ 1 + (y * v) / (n+1)) 1 49 | let err : ℚ := 2 * y^(n+1) / (n+1).factorial 50 | (ey + err) ^ xc 51 | 52 | lemma List_foldr_eq_finset_sum (x : ℚ) (n : ℕ) : 53 | List.foldr (fun n v => 1 + x * v / (↑n + 1)) 1 (List.range n) 54 | = ∑ i ∈ Finset.range (n + 1), x^i / i.factorial := by 55 | suffices ∀ (v : ℚ), List.foldr (fun n v => 1 + x * v / (↑n + 1)) v (List.range n) 56 | = (∑ i ∈ Finset.range n, x^i / i.factorial) + v * (x ^ n) / n.factorial by 57 | convert this 1 58 | simp [Finset.range_add_one, _root_.add_comm] 59 | induction n 60 | · simp 61 | · intro v 62 | rename_i n ih 63 | simp [Finset.range_add_one, List.range_succ, Nat.factorial_succ, ih] 64 | rw [add_mul, one_mul, add_div, pow_succ] 65 | suffices x * v / (↑n + 1) * x ^ n / ↑n.factorial = v * (x ^ n * x) / ((↑n + 1) * ↑n.factorial) by 66 | rw [this] 67 | ring 68 | field_simp 69 | ring_nf 70 | 71 | theorem exp_lb₀_pos {x : ℚ} (n : ℕ) (hx : 0 ≤ x) : 0 < exp_lb₀ x n := by 72 | rw [exp_lb₀, List_foldr_eq_finset_sum, Finset.range_add_one'] 73 | rw [Finset.sum_insert (by simp)] 74 | positivity 75 | 76 | theorem exp_lb₀_ge_one {x : ℚ} (n : ℕ) (hx : 0 ≤ x) : 1 ≤ exp_lb₀ x n := by 77 | rw [exp_lb₀, List_foldr_eq_finset_sum, Finset.range_add_one'] 78 | rw [Finset.sum_insert (by simp)] 79 | apply one_le_pow₀ 80 | simpa using by positivity 81 | 82 | theorem exp_lb₀_le_exp {x : ℚ} (n : ℕ) (hx : 0 ≤ x) : exp_lb₀ x n ≤ Real.exp x := by 83 | rw [exp_lb₀, List_foldr_eq_finset_sum] 84 | have he : Real.exp x = (Real.exp (x / ↑⌈x⌉.toNat)) ^ ⌈x⌉.toNat := by 85 | cases eq_or_lt_of_le hx 86 | · subst x; simp 87 | · rw [div_eq_mul_inv, Real.exp_mul, Real.rpow_inv_natCast_pow (Real.exp_nonneg ↑x)] 88 | rwa [Nat.ne_zero_iff_zero_lt, Int.lt_toNat, Int.lt_ceil] 89 | rw [he] 90 | push_cast 91 | apply pow_le_pow_left₀ (by positivity) 92 | apply_mod_cast Real.sum_le_exp_of_nonneg 93 | positivity 94 | 95 | theorem exp_ub₀_ge_exp (x : ℚ) (n : ℕ) (hx : 0 ≤ x) : Real.exp x ≤ exp_ub₀ x n := by 96 | rw [exp_ub₀, List_foldr_eq_finset_sum] 97 | have he : Real.exp x = (Real.exp (x / ↑⌈x⌉.toNat)) ^ ⌈x⌉.toNat := by 98 | cases eq_or_lt_of_le hx 99 | · subst x; simp 100 | · rw [div_eq_mul_inv, Real.exp_mul, Real.rpow_inv_natCast_pow (Real.exp_nonneg ↑x)] 101 | rwa [Nat.ne_zero_iff_zero_lt, Int.lt_toNat, Int.lt_ceil] 102 | rw [he] 103 | push_cast 104 | apply pow_le_pow_left₀ (by positivity) 105 | have := Real.exp_bound' (x := (↑x / ↑⌈x⌉.toNat)) (by positivity) ?_ (Nat.zero_lt_succ n) 106 | · refine le_trans this ?_ 107 | simp only [Nat.succ_eq_add_one, add_le_add_iff_left] 108 | rw [_root_.mul_comm 2, ← mul_div, ← mul_div] 109 | apply mul_le_mul_of_nonneg_left ?_ (by positivity) 110 | rw [_root_.mul_comm, ← div_div] 111 | apply div_le_div_of_nonneg_right ?_ (by positivity) 112 | rw [add_div, div_self (by positivity), ← one_add_one_eq_two, add_le_add_iff_left 1, one_div] 113 | apply inv_le_one_of_one_le₀ 114 | simp 115 | · cases eq_or_lt_of_le hx 116 | · subst x; simp 117 | rw [div_le_one₀] 118 | · trans (⌈x⌉ : ℝ) 119 | · simpa using Int.le_ceil x 120 | · norm_cast 121 | simp only [Int.ofNat_toNat, le_sup_left] 122 | simpa 123 | 124 | -- theorem _root_.Real.mul_log_one_plus_div_lt {x t : ℝ} (hx : 0 < x) (ht : 0 < t) : 125 | -- x * Real.log (1 + t / x) < t := by 126 | -- suffices Real.log (1 + t / x) < t / x by 127 | -- rwa [_root_.mul_comm, ← lt_div_iff₀ hx] 128 | -- convert Real.log_lt_sub_one_of_pos (x := 1 + t / x) (by positivity) ?_ using 1 129 | -- · ring_nf 130 | -- · simpa using ⟨ht.ne', hx.ne'⟩ 131 | 132 | -- theorem _root_.Real.mul_log_one_plus_div_le {x t : ℝ} (hx : 0 < x) (ht : 0 ≤ t) : 133 | -- x * Real.log (1 + t / x) ≤ t := by 134 | -- cases eq_or_lt_of_le ht 135 | -- · subst t; simp 136 | -- · apply (Real.mul_log_one_plus_div_lt hx ‹_›).le 137 | 138 | -- theorem _root_.Real.one_plus_div_pow_lt_exp {x t : ℝ} (hx : 0 < x) (ht : 0 < t) : 139 | -- (1 + t / x) ^ x < Real.exp t := by 140 | -- convert Real.exp_lt_exp_of_lt (x := x * Real.log (1 + t / x)) 141 | -- (Real.mul_log_one_plus_div_lt hx ht) using 1 142 | -- rw [← Real.log_rpow (by positivity), Real.exp_log (by positivity)] 143 | 144 | theorem _root_.Real.one_plus_pow_lt_exp_mul {x y : ℝ} (hx : 0 < x) (hy : 0 < y) : 145 | (1 + x) ^ y < Real.exp (x * y) := by 146 | rw [Real.exp_mul, _root_.add_comm] 147 | exact Real.rpow_lt_rpow (by linarith) (Real.add_one_lt_exp hx.ne') hy 148 | 149 | /-- This proves that the gap between `exp_lb₀` and `exp_ub₀` shrinks at least as 150 | fast as `exp x * (exp (2x/n!) - 1)`. This is then used to prove that they are 151 | cauchy sequences in `n`, i.e. taking sufficiently large `n` makes this go to zero. -/ 152 | theorem exp_ub₀_sub_exp_lb₀ {x : ℚ} (n : ℕ) (hx : 0 ≤ x) : 153 | exp_ub₀ x n - exp_lb₀ x n ≤ Real.exp x * (Real.exp (2 * x / n.factorial) - 1) := by 154 | rw [exp_ub₀, exp_lb₀] 155 | rw [List_foldr_eq_finset_sum] 156 | 157 | --Special case out x=0 158 | rcases eq_or_lt_of_le hx 159 | · subst x 160 | simp 161 | clear hx; rename_i hx 162 | 163 | set y := ∑ i ∈ Finset.range (n + 1), (x / ↑⌈x⌉.toNat) ^ i / ↑i.factorial with hdy 164 | set z := 2 * (x / ↑⌈x⌉.toNat) ^ (n + 1) / ↑(n + 1).factorial 165 | 166 | 167 | have hy : y ≠ 0 := by 168 | apply ne_of_gt 169 | rw [hdy, Finset.range_add_one', Finset.sum_insert (by simp)] 170 | positivity 171 | conv_lhs => 172 | equals (y ^ ⌈x⌉.toNat * ((1 + z / y)^⌈x⌉.toNat - 1) : ℝ) => 173 | rw [mul_sub, mul_one, ← mul_pow, mul_add, mul_one, mul_div_cancel₀] 174 | · norm_cast 175 | · norm_cast 176 | 177 | have hxn : 0 < ↑⌈x⌉.toNat := by simpa 178 | have hz : 0 < z := by positivity 179 | have hzy₀ : 0 < z / y := by positivity 180 | have hy₂ : y ^ ⌈x⌉.toNat ≤ Real.exp x := by 181 | have := exp_lb₀_le_exp n hx.le 182 | rw [exp_lb₀, List_foldr_eq_finset_sum] at this 183 | exact_mod_cast this 184 | 185 | have hzy₁ : z / y ≤ 2 * (x / ↑⌈x⌉.toNat) / n.factorial := by 186 | trans z 187 | · apply div_le_self hz.le 188 | rw [hdy, Finset.range_add_one'] 189 | simp 190 | positivity 191 | · unfold z 192 | refine div_le_div₀ (by positivity) ?_ (by positivity) ?_ 193 | · simp only [pow_succ, Nat.ofNat_pos, mul_le_mul_left] 194 | refine mul_le_of_le_one_left (by positivity) ?_ 195 | apply pow_le_one₀ (by positivity) 196 | rw [div_le_one₀ (by positivity)] 197 | refine (Int.le_ceil x).trans ?_ 198 | exact_mod_cast Int.self_le_toNat ⌈x⌉ 199 | · exact_mod_cast Nat.factorial_le (Nat.le_add_right n 1) 200 | 201 | have hzy₂ := calc ((1 + z / y) ^ ⌈x⌉.toNat : ℝ) 202 | _ ≤ (1 + 2 * (x / ⌈x⌉.toNat) / n.factorial) ^ ⌈x⌉.toNat := by 203 | norm_cast at hzy₁ ⊢ 204 | apply pow_le_pow_left₀ (by positivity) (by linarith) 205 | _ < Real.exp (2 * (x / ⌈x⌉.toNat) * ⌈x⌉.toNat / n.factorial) := by 206 | convert Real.one_plus_pow_lt_exp_mul (y := ⌈x⌉.toNat) 207 | (Rat.cast_pos.mpr <| lt_of_lt_of_le hzy₀ hzy₁) (by simpa) using 2 208 | · norm_cast 209 | · norm_cast 210 | ring_nf 211 | _ ≤ Real.exp (2 * x / n.factorial) := by 212 | rw [mul_div, div_mul, div_self, div_one] 213 | simpa 214 | 215 | apply mul_le_mul hy₂ (by simpa using hzy₂.le) ?_ (Real.exp_nonneg _) 216 | rw [sub_nonneg] 217 | apply one_le_pow₀ 218 | rw [le_add_iff_nonneg_right] 219 | positivity 220 | 221 | /-- Unlike `exp_lb₀`, which only works for `0 ≤ x`, this is valid for all `x`. -/ 222 | def exp_lb (x : ℚ) : ℕ → ℚ := 223 | if 0 ≤ x then exp_lb₀ x else fun n ↦ (exp_ub₀ (-x) n)⁻¹ 224 | 225 | /-- Unlike `exp_ub₀`, which only works for `0 ≤ x`, this is valid for all `x`. -/ 226 | def exp_ub (x : ℚ) : ℕ → ℚ := 227 | if 0 ≤ x then exp_ub₀ x else fun n ↦ (exp_lb₀ (-x) n)⁻¹ 228 | 229 | theorem exp_lb_le_exp (x : ℚ) (n : ℕ) : exp_lb x n ≤ Real.exp x := by 230 | rw [exp_lb] 231 | split 232 | · apply exp_lb₀_le_exp n ‹_› 233 | · simp only [Rat.cast_inv] 234 | rw [← inv_inv (Real.exp _)] 235 | have := exp_ub₀_ge_exp (-x) n (by linarith) 236 | rw [inv_le_inv₀] 237 | · simpa [Real.exp_neg] 238 | · refine lt_of_lt_of_le ?_ this 239 | positivity 240 | · positivity 241 | 242 | theorem exp_ub_ge_exp (x : ℚ) (n : ℕ) : Real.exp x ≤ exp_ub x n := by 243 | rw [exp_ub] 244 | split 245 | · apply exp_ub₀_ge_exp x n ‹_› 246 | · simp only [Rat.cast_inv] 247 | rw [← inv_inv (Real.exp _)] 248 | have := exp_lb₀_le_exp (x := -x) n (by linarith) 249 | rw [inv_le_inv₀] 250 | · simpa [Real.exp_neg] 251 | · positivity 252 | · exact_mod_cast exp_lb₀_pos (x := -x) n (by linarith) 253 | 254 | theorem exp_ub_sub_exp_lb_of_nonneg {x : ℚ} (n : ℕ) (hx : 0 ≤ x) : 255 | exp_ub x n - exp_lb x n ≤ Real.exp x * (Real.exp (2 * x / n.factorial) - 1) := by 256 | simpa [exp_ub, hx, ↓reduceIte, exp_lb] using exp_ub₀_sub_exp_lb₀ n hx 257 | 258 | theorem exp_ub_sub_exp_lb_of_neg {x : ℚ} (n : ℕ) (hx : x < 0) : 259 | exp_ub x n - exp_lb x n ≤ Real.exp (2 * ↑(-x) / n.factorial) - 1 := by 260 | simp only [exp_ub, Rat.not_le.mpr hx, ↓reduceIte, exp_lb] 261 | replace hx : 0 < -x := by linarith 262 | generalize -x=x' at hx ⊢; clear x; rename ℚ => x 263 | 264 | have hl₁ := exp_lb₀_pos n hx.le 265 | have hl₂ := exp_lb₀_le_exp n hx.le 266 | have hu₁ := exp_ub₀_ge_exp x n hx.le 267 | have hu₂ : 0 < exp_ub₀ x n := by rify at hl₁ ⊢; linarith 268 | have hlu := exp_ub₀_sub_exp_lb₀ n hx.le 269 | 270 | conv_lhs => 271 | equals (exp_ub₀ x n - exp_lb₀ x n : ℝ) / (exp_ub₀ x n * exp_lb₀ x n) => 272 | field_simp 273 | left 274 | ring_nf 275 | 276 | rw [div_le_iff₀ (by positivity)] 277 | refine hlu.trans ?_ 278 | conv_rhs => 279 | rw [_root_.mul_comm (Rat.cast _), ← _root_.mul_assoc, _root_.mul_comm] 280 | 281 | apply mul_le_mul hu₁ ?_ ?_ (by positivity) 282 | · apply le_mul_of_one_le_right 283 | · rw [sub_nonneg, Real.one_le_exp_iff] 284 | positivity 285 | · exact_mod_cast exp_lb₀_ge_one n hx.le 286 | · rw [sub_nonneg, Real.one_le_exp_iff] 287 | positivity 288 | 289 | theorem exp_ub_lb_err (x : ℚ) (n : ℕ) : 290 | Real.exp x ≤ exp_lb x n + (Real.exp ↑|x| * (Real.exp (2 * |x| / n.factorial) - 1)) ∧ 291 | exp_ub x n - (Real.exp ↑|x| * (Real.exp (2 * |x| / n.factorial) - 1)) ≤ Real.exp x := by 292 | have hl := exp_lb_le_exp x n 293 | have hu := exp_ub_ge_exp x n 294 | rcases lt_or_ge x 0 with hx | hx 295 | · have hlu := exp_ub_sub_exp_lb_of_neg n hx 296 | replace hlu := hlu.trans (le_mul_of_one_le_left (a := Real.exp ↑(-x)) ?_ ?_) 297 | · simp only [abs_of_neg hx] 298 | constructor <;> linarith 299 | · rw [sub_nonneg, Real.one_le_exp_iff] 300 | have : 0 < -x := by linarith 301 | positivity 302 | · simp [hx.le] 303 | · have hlu := exp_ub_sub_exp_lb_of_nonneg n hx 304 | simp only [abs_of_nonneg hx] 305 | constructor <;> linarith 306 | 307 | private lemma err_antitone_n (x : ℝ) : 308 | Antitone fun (n : ℕ) ↦ Real.exp ↑|x| * (Real.exp (2 * |x| / n.factorial) - 1) := by 309 | apply Antitone.const_mul ?_ (Real.exp_nonneg |x|) 310 | simp_rw [sub_eq_add_neg] 311 | apply Antitone.add_const 312 | conv_rhs => 313 | equals ((Real.exp) ∘ (fun n ↦ (2 * |x| / n.factorial))) => 314 | ext x 315 | rfl 316 | refine Monotone.comp_antitone Real.exp_monotone ?_ 317 | simp_rw [div_eq_mul_inv] 318 | apply Antitone.const_mul ?_ (by positivity) 319 | conv_rhs => 320 | equals ((fun (x : ℕ) ↦ (↑x)⁻¹) ∘ (fun n ↦ n.factorial)) => 321 | ext x 322 | rfl 323 | intro x y h 324 | field_simp 325 | apply one_div_le_one_div_of_le 326 | · exact Nat.cast_pos'.mpr (Nat.factorial_pos x) 327 | · exact Nat.cast_le.mpr (Nat.factorial_le h) 328 | 329 | private lemma err_monotone_x (n : ℕ) : 330 | MonotoneOn (fun (x : ℝ) ↦ Real.exp x * (Real.exp (2 * x / n.factorial) - 1)) (Set.Ici 0) := by 331 | apply MonotoneOn.mul (Real.exp_monotone.monotoneOn _) ?_ (fun _ ↦ Real.exp_nonneg ·) ?_ 332 | · simp_rw [sub_eq_add_neg] 333 | apply Monotone.monotoneOn 334 | apply Monotone.add_const 335 | conv_rhs => 336 | equals ((Real.exp) ∘ (fun x ↦ (2 * x / n.factorial))) => 337 | ext x 338 | rfl 339 | refine Monotone.comp Real.exp_monotone ?_ 340 | simp_rw [_root_.mul_comm, ← mul_div] 341 | exact Monotone.mul_const monotone_id (by positivity) 342 | · intro x (hx : 0 ≤ x) 343 | rw [sub_nonneg, Real.one_le_exp_iff] 344 | positivity 345 | 346 | private lemma inverr_monotone_x (ε : ℝ) (hε : 0 < ε): 347 | MonotoneOn (fun (x : ℝ) ↦ 2 * x / Real.log (1 + ε / Real.exp x)) (Set.Ici 0) := by 348 | apply MonotoneOn.mul ?_ ?_ ?_ ?_ 349 | · apply MonotoneOn.mul monotoneOn_const ?_ (by norm_num) (by norm_num) 350 | intro x y h 351 | exact fun x_1 a => a 352 | · conv_lhs => 353 | equals ((·⁻¹) ∘ Real.log ∘ (1 + ε / Real.exp ·)) => 354 | ext x 355 | rfl 356 | suffices AntitoneOn (1 + ε / Real.exp ·) (Set.Ici 0) by 357 | intro x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy 358 | specialize this hx hy hxy 359 | dsimp at this ⊢ 360 | replace this := Real.log_le_log (by positivity) this 361 | rwa [inv_le_inv₀] 362 | all_goals ( 363 | apply Real.log_pos 364 | rw [lt_add_iff_pos_right] 365 | positivity 366 | ) 367 | apply AntitoneOn.const_add 368 | intro x (hx : 0 ≤ x) y (hy : 0 ≤ y) hxy 369 | refine div_le_div₀ hε.le le_rfl (Real.exp_pos _) (Real.exp_le_exp.mpr hxy) 370 | · intro x (h : 0 ≤ x) 371 | positivity 372 | · intros 373 | rw [inv_nonneg, Real.log_nonneg_iff] 374 | · rw [le_add_iff_nonneg_right]; positivity 375 | · positivity 376 | 377 | private lemma exists_n_bound_err (a b : ℝ) {ε : ℝ} (hε : 0 < ε) : 378 | ∃ (n : ℕ), ∀ (y : ℚ), ↑y ∈ Set.Icc a b → 379 | (Real.exp ↑|y| * (Real.exp (2 * |y| / n.factorial) - 1)) ≤ ε 380 | := by 381 | conv => 382 | enter [1,n,y,h] 383 | equals 2 * |y| / Real.log (1 + ε / Real.exp |y|) ≤ n.factorial => 384 | by_cases hy : y = 0 385 | · simp [hε.le, hy] 386 | have hy₂ : 0 < |y| := by simpa 387 | rw [← Rat.cast_abs] 388 | rw [_root_.mul_comm, ← le_div_iff₀ (by positivity)] 389 | rw [tsub_le_iff_right, _root_.add_comm] 390 | rw [← Real.le_log_iff_exp_le (by positivity)] 391 | rw [eq_iff_iff] 392 | apply div_le_comm₀ (by positivity) (Real.log_pos ?_) 393 | rw [lt_add_iff_pos_right] 394 | positivity 395 | let v₁ := 2 * |a| / Real.log (1 + ε / Real.exp |a|) 396 | let v₂ := 2 * |b| / Real.log (1 + ε / Real.exp |b|) 397 | let v₃ := ⌈max v₁ v₂⌉.toNat 398 | use v₃ 399 | intro y ⟨hy₁, hy₂⟩ 400 | trans (v₃ : ℝ) 401 | · trans max v₁ v₂ 402 | · trans 2 * (|a| ⊔ |b|) / Real.log (1 + ε / Real.exp (|a| ⊔ |b|)) 403 | · rw [Rat.cast_abs] 404 | exact inverr_monotone_x ε hε (abs_nonneg (y : ℝ)) (show 0 ≤ |a| ⊔ |b| by positivity) 405 | (abs_le_max_abs_abs hy₁ hy₂) 406 | · apply le_of_eq 407 | apply MonotoneOn.map_max (inverr_monotone_x ε hε) 408 | · exact abs_nonneg a 409 | · exact abs_nonneg b 410 | · unfold v₃ 411 | trans ↑⌈v₁ ⊔ v₂⌉ 412 | · exact Int.le_ceil (v₁ ⊔ v₂) 413 | · norm_cast 414 | simp 415 | · rw [Nat.cast_le] 416 | exact Nat.self_le_factorial v₃ 417 | 418 | theorem TLUW_lb_ub : 419 | TendstoLocallyUniformlyWithout (fun n x => exp_lb x n) Real.exp ∧ 420 | TendstoLocallyUniformlyWithout (fun n x => exp_ub x n) Real.exp := by 421 | rw [TendstoLocallyUniformlyWithout] 422 | constructor 423 | all_goals ( 424 | intro ε hε x 425 | obtain ⟨n,hn⟩ := exists_n_bound_err (x-1) (x+1) (half_pos hε) 426 | use Set.Icc (x-1) (x+1), Icc_mem_nhds (sub_one_lt x) (lt_add_one x), n 427 | intro b hb y ⟨hy₁, hy₂⟩ 428 | specialize hn y ⟨hy₁, hy₂⟩ 429 | rw [abs_sub_lt_iff] 430 | have hl := exp_lb_le_exp y b 431 | have hu := exp_ub_ge_exp y b 432 | have ⟨h₀, h₁⟩ := exp_ub_lb_err y b 433 | have h₂ := err_antitone_n y hb 434 | simp only [Rat.cast_abs] at h₀ h₁ hn 435 | constructor <;> linarith 436 | ) 437 | 438 | end Exp 439 | 440 | def exp : ComputableℝSeq → ComputableℝSeq := 441 | of_TendstoLocallyUniformly_Continuous Real.continuous_exp 442 | (fun n q ↦ ⟨⟨Exp.exp_lb q.fst n, Exp.exp_ub q.snd n⟩, 443 | Rat.cast_le.mp <| 444 | le_trans (Exp.exp_lb_le_exp q.fst n) <| 445 | le_trans (Real.exp_le_exp_of_le <| Rat.cast_le.mpr q.fst_le_snd) <| 446 | Exp.exp_ub_ge_exp q.snd n⟩) 447 | (fun n x ↦ Exp.exp_lb x n) 448 | (fun n x ↦ Exp.exp_ub x n) 449 | (fun n ⟨⟨q₁, _⟩, _⟩ _ ⟨hx, _⟩ ↦ 450 | (Exp.exp_lb_le_exp q₁ n).trans (Real.exp_le_exp_of_le hx)) 451 | (fun n ⟨⟨_, q₂⟩, _⟩ _ ⟨_, hx⟩ ↦ 452 | (Real.exp_le_exp_of_le hx).trans (Exp.exp_ub_ge_exp q₂ n)) 453 | (fun _ _ ↦ rfl) 454 | (Exp.TLUW_lb_ub.1) 455 | (Exp.TLUW_lb_ub.2) 456 | 457 | end ComputableℝSeq 458 | 459 | namespace IsComputable 460 | 461 | instance instComputableExp (x : ℝ) [hx : IsComputable x] : IsComputable (Real.exp x) := 462 | lift Real.exp ComputableℝSeq.exp 463 | (by apply ComputableℝSeq.val_of_TendstoLocallyUniformly_Continuous) hx 464 | 465 | instance instComputableSinh (x : ℝ) [hx : IsComputable x] : IsComputable (Real.sinh x) := 466 | lift_eq (Real.sinh_eq x).symm inferInstance 467 | 468 | instance instComputableCosh (x : ℝ) [hx : IsComputable x] : IsComputable (Real.cosh x) := 469 | lift_eq (Real.cosh_eq x).symm inferInstance 470 | 471 | instance instComputableTanh (x : ℝ) [hx : IsComputable x] : IsComputable (Real.tanh x) := 472 | lift_eq (Real.tanh_eq_sinh_div_cosh x).symm inferInstance 473 | 474 | end IsComputable 475 | -------------------------------------------------------------------------------- /ComputableReal/SpecialFunctions/Pi.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.SpecialFunctions.Sqrt 2 | import Mathlib.Data.Real.Pi.Bounds 3 | import Mathlib.Analysis.SpecialFunctions.Trigonometric.Basic 4 | 5 | open scoped QInterval 6 | 7 | namespace ComputableℝSeq 8 | 9 | section Pi 10 | 11 | instance instComputableSqrtTwoAddSeries (x : ℝ) [hx : IsComputable x] (n : ℕ) : 12 | IsComputable (Real.sqrtTwoAddSeries x n) := 13 | n.rec hx (fun _ _ ↦ IsComputable.instComputableSqrt _) 14 | 15 | def sqrtTwoAddSeries_n : ℕ → ComputableℝSeq := 16 | fun n ↦ (instComputableSqrtTwoAddSeries 0 n).seq 17 | 18 | theorem sqrtTwoAddSeries_n_lb_le (n k : ℕ) : (sqrtTwoAddSeries_n n).lb k ≤ Real.sqrtTwoAddSeries 0 n := by 19 | rw [← IsComputable.prop (x := Real.sqrtTwoAddSeries 0 n)] 20 | exact ComputableℝSeq.hlb _ _ 21 | 22 | theorem sqrtTwoAddSeries_n_ub_ge (n k : ℕ) : Real.sqrtTwoAddSeries 0 n ≤ (sqrtTwoAddSeries_n n).ub k := by 23 | rw [← IsComputable.prop (x := Real.sqrtTwoAddSeries 0 n)] 24 | exact ComputableℝSeq.hub _ _ 25 | 26 | theorem sqrtTwoAddSeries_n_ub_pos (n k : ℕ) : (0 : ℝ) ≤ (sqrtTwoAddSeries_n n).ub k := by 27 | exact le_trans (Real.sqrtTwoAddSeries_zero_nonneg n) (sqrtTwoAddSeries_n_ub_ge n k) 28 | 29 | theorem sqrtTwoAddSeries_n_lb_lt_two (n k : ℕ) : (sqrtTwoAddSeries_n n).lb k < 2 := by 30 | rify 31 | refine lt_of_le_of_lt ?_ (Real.sqrtTwoAddSeries_lt_two n) 32 | convert ComputableℝSeq.hlb _ _ 33 | symm 34 | exact IsComputable.prop 35 | 36 | theorem sqrtTwoAddSeries_n_succ_lb (n k : ℕ) : 37 | (sqrtTwoAddSeries_n (n + 1)).lb k = (Sqrt.sqrtq (2 + (sqrtTwoAddSeries_n n).lub k) k).fst := by 38 | rfl 39 | 40 | theorem sqrtTwoAddSeries_n_succ_ub (n k : ℕ) : 41 | (sqrtTwoAddSeries_n (n + 1)).ub k = (Sqrt.sqrtq (2 + (sqrtTwoAddSeries_n n).lub k) k).snd := by 42 | rfl 43 | 44 | theorem sqrtTwoAddSeries_n_lb_nonneg (n k : ℕ) : 0 ≤ (sqrtTwoAddSeries_n n).lb k := by 45 | cases n 46 | · rfl 47 | · apply Sqrt.sqrtq_nonneg 48 | 49 | theorem sqrtTwoAddSeries_n_lb_gt_one (n k : ℕ) (hk : 3 ≤ k) : 1 ≤ (sqrtTwoAddSeries_n (n + 1)).lb k := by 50 | have h₀ : (0 : ℝ) ≤ ((sqrtTwoAddSeries_n n).lub k).snd := sqrtTwoAddSeries_n_ub_pos n k 51 | rw [sqrtTwoAddSeries_n_succ_lb, Sqrt.sqrt_lb_def, 52 | if_neg (by push_neg; change 0 < 2 + _; rify; positivity)] 53 | clear h₀ 54 | 55 | have h₁ := sqrtTwoAddSeries_n_lb_nonneg n k 56 | rify at h₁ ⊢ 57 | 58 | rw [show (2 + (sqrtTwoAddSeries_n n).lub k).toProd.1 = 2 + (sqrtTwoAddSeries_n n).lb k by rfl] 59 | set x : ℚ := (sqrtTwoAddSeries_n n).lb k 60 | have h₂ := Sqrt.sqrt_le_mkRat_add (2 + x) k 61 | generalize ↑(mkRat (Int.sqrt ((2 + x).num * 4 ^ k)) (((2 + x).den * 4 ^ k).sqrt + 1))=y at h₂ ⊢ 62 | replace h₂ : √↑(2 + x) * (1 - 2 / 2^k) ≤ y := by 63 | ring_nf at h₂ ⊢ 64 | linarith 65 | 66 | refine le_trans ?_ (le_trans (a := √↑(2 + x) * 3 / 4) ?_ h₂) 67 | · suffices 4 / 3 ≤ √↑(2 + x) by linarith 68 | apply Real.le_sqrt_of_sq_le 69 | push_cast 70 | linarith 71 | · rw [← mul_div] 72 | apply mul_le_mul_of_nonneg_left ?_ (Real.sqrt_nonneg _) 73 | rw [show (3 / 4 : ℝ) = 1 - 2 / 8 by norm_num] 74 | apply sub_le_sub_left 75 | apply div_le_div₀ zero_le_two le_rfl Nat.ofNat_pos' 76 | rw [show 8 = (2 : ℝ) ^ 3 by norm_num] 77 | exact_mod_cast Nat.pow_le_pow_of_le_right Nat.ofNat_pos' hk 78 | 79 | theorem sqrtTwoAddSeries_n_bounds (n k : ℕ) (hk : 3 ≤ k) : 80 | (sqrtTwoAddSeries_n n).ub k ≤ (sqrtTwoAddSeries_n n).lb k + 18 * n / 2^k 81 | := by 82 | induction n 83 | · simp only [CharP.cast_eq_zero, zero_div, mul_zero, add_zero] 84 | rfl 85 | rename_i n ih 86 | dsimp [lb, ub] at ih 87 | rify at ih ⊢ 88 | rw [sqrtTwoAddSeries_n_succ_lb, sqrtTwoAddSeries_n_succ_ub] 89 | 90 | set x := ((sqrtTwoAddSeries_n n).lub k) 91 | 92 | have hl : (2 + x).fst ≤ 2 + Real.sqrtTwoAddSeries 0 n := by 93 | change ((2 + _ : ℚ) : ℝ) ≤ _ 94 | rw [Rat.cast_add, Rat.cast_ofNat, add_le_add_iff_left] 95 | convert ComputableℝSeq.hlb _ _ 96 | symm 97 | exact IsComputable.prop 98 | 99 | have hu : 2 + Real.sqrtTwoAddSeries 0 n ≤ (2 + x).snd := by 100 | change _ ≤ ((2 + _ : ℚ) : ℝ) 101 | rw [Rat.cast_add, Rat.cast_ofNat, add_le_add_iff_left] 102 | rw [← ge_iff_le] 103 | convert ComputableℝSeq.hub _ _ 104 | symm 105 | exact IsComputable.prop 106 | 107 | have hm : 0 < (2 + x).fst := by 108 | change 0 < 2 + _ 109 | suffices 0 ≤ x.fst by linarith 110 | apply sqrtTwoAddSeries_n_lb_nonneg n k 111 | 112 | have h₁ := Sqrt.sqrt_le_sqrtq_add (2 + Real.sqrtTwoAddSeries 0 n) (2 + x) k ⟨hl, hu⟩ hm 113 | have h₂ := Sqrt.sqrtq_sub_le_sqrt (2 + Real.sqrtTwoAddSeries 0 n) (2 + x) k ⟨hl, hu⟩ hm hk 114 | 115 | simp only [show (2 + x).snd = 2 + x.snd by rfl, show (2 + x).fst = 2 + x.fst by rfl] at h₁ h₂ 116 | simp only [Rat.cast_add, Rat.cast_ofNat, add_sub_add_left_eq_sub, tsub_le_iff_right] at h₁ h₂ 117 | 118 | set y := √(2 + Real.sqrtTwoAddSeries 0 n) 119 | set z := (Sqrt.sqrtq (2 + x) k).fst 120 | set w := (Sqrt.sqrtq (2 + x) k).snd 121 | set x₁ := x.fst 122 | set x₂ := x.snd 123 | 124 | have hgap : w - z ≤ (3 / 2) * ((↑x₂ - ↑x₁) / (√(2 + ↑x₁))) + 9 * y / 2^k := by 125 | rw [← mul_div] at h₁ h₂ ⊢ 126 | conv at h₁ => 127 | enter [2,2,2] 128 | rw [_root_.mul_comm] 129 | rw [← div_div] at h₁ 130 | linarith 131 | 132 | have hxdiv : 3 / 2 * ((↑x₂ - ↑x₁) / √(2 + ↑x₁)) ≤ 18 * ↑n / 2 ^ k := by 133 | rcases n with _|n 134 | · simp [show x₁ = 0 by rfl, show x₂ = 0 by rfl] 135 | suffices (3 / 2) / √(2 + ↑x₁) ≤ 1 by 136 | rw [mul_div, _root_.mul_comm, ← mul_div] 137 | apply mul_le_of_le_of_le_one_of_nonneg 138 | · linarith 139 | · exact this 140 | · have : (x.fst : ℝ) ≤ x.snd := by exact_mod_cast x.fst_le_snd 141 | linarith [x.fst_le_snd] 142 | have hx₂ : 1 ≤ x₁ := 143 | sqrtTwoAddSeries_n_lb_gt_one _ _ hk 144 | rify at hx₂ 145 | rw [div_le_one₀ (by positivity), Real.le_sqrt' (by positivity)] 146 | linarith 147 | 148 | have hydiv : 9 * y / 2^k ≤ 18 / 2^k := by 149 | rw [div_le_div_iff_of_pos_right] 150 | · have hy2 : y < 2 := Real.sqrtTwoAddSeries_lt_two (n.succ) 151 | linarith 152 | · exact_mod_cast Nat.two_pow_pos k 153 | 154 | rw [mul_add, add_div] 155 | linarith 156 | 157 | def sqrtTwoSubSqrtTwoAddSeries_n : ℕ → ComputableℝSeq := 158 | fun n ↦ (inferInstance : IsComputable (Real.sqrt (2 - Real.sqrtTwoAddSeries 0 n))).seq 159 | 160 | theorem sqrtTwoSubSqrtTwoAddSeries_eq (n k : ℕ) : 161 | (sqrtTwoSubSqrtTwoAddSeries_n n).lub k = Sqrt.sqrtq ((2 - sqrtTwoAddSeries_n n).lub k) k := by 162 | rfl 163 | 164 | theorem real_sqrtTwoAddSeries_lb (n : ℕ) : 1 / 2 ^ n < √(2 - Real.sqrtTwoAddSeries 0 n) := by 165 | have h : 2 ≤ Real.pi - 1 / 4 ^ n := by 166 | have : 1 / 4 ^ n ≤ (1 : ℝ) := by 167 | rw [one_div] 168 | apply inv_le_one_of_one_le₀ 169 | exact_mod_cast Nat.one_le_pow' n 3 170 | linarith [Real.pi_gt_three] 171 | 172 | have h₂ := lt_of_le_of_lt h (sub_right_lt_of_lt_add (Real.pi_lt_sqrtTwoAddSeries n)) 173 | ring_nf at h₂ 174 | rwa [lt_mul_iff_one_lt_left zero_lt_two, ← div_lt_iff₀ (by positivity)] at h₂ 175 | 176 | theorem sqrtTwoSubSqrtTwoAddSeries_lb (n k : ℕ) (hk : 3 ≤ k) : 177 | √(2 - Real.sqrtTwoAddSeries 0 n) ≤ (sqrtTwoSubSqrtTwoAddSeries_n n).lb k + 178 | (18 * n * 2 ^ n + 4) / 2 ^ k 179 | := by 180 | 181 | dsimp [lb] 182 | rw [sqrtTwoSubSqrtTwoAddSeries_eq] 183 | 184 | have h₁ := Sqrt.sqrt_le_sqrtq_add' (2 - Real.sqrtTwoAddSeries 0 n) ((2 - sqrtTwoAddSeries_n n).lub k) k 185 | ⟨?_, ?_⟩ ?_; rotate_left 186 | · change ((2 + -_ : ℚ) : ℝ) ≤ _ 187 | rw [Rat.cast_add, Rat.cast_ofNat, Rat.cast_neg] 188 | apply tsub_le_tsub_left 189 | rw [← ge_iff_le] 190 | convert ComputableℝSeq.hub _ _ 191 | symm 192 | exact IsComputable.prop 193 | · change _ ≤ ((2 + -_ : ℚ) : ℝ) 194 | rw [Rat.cast_add, Rat.cast_ofNat, Rat.cast_neg] 195 | apply tsub_le_tsub_left 196 | convert ComputableℝSeq.hlb _ _ 197 | symm 198 | exact IsComputable.prop 199 | · linarith [Real.sqrtTwoAddSeries_lt_two n] 200 | 201 | have hx : ((2 - sqrtTwoAddSeries_n n).lub k).fst = 2 - (sqrtTwoAddSeries_n n).ub k := by 202 | change (2 +- _) = _ 203 | rw [sub_eq_add_neg] 204 | rfl 205 | have hy : ((2 - sqrtTwoAddSeries_n n).lub k).snd = 2 - (sqrtTwoAddSeries_n n).lb k := by 206 | change (2 +- _) = _ 207 | rw [sub_eq_add_neg] 208 | rfl 209 | simp only [hx, hy, Rat.cast_sub, Rat.cast_ofNat, sub_sub_sub_cancel_left] at h₁ 210 | clear hx hy 211 | 212 | have h₄ := Real.sqrtTwoAddSeries_zero_nonneg n 213 | have h₅ := Real.sqrtTwoAddSeries_lt_two n 214 | have h₆ := sqrtTwoAddSeries_n_bounds n k hk 215 | have h₇ := (real_sqrtTwoAddSeries_lb n).le 216 | 217 | generalize 218 | (Sqrt.sqrtq ((2 - sqrtTwoAddSeries_n n).lub k) k).fst=w, 219 | (sqrtTwoAddSeries_n n).lb k=x, 220 | (sqrtTwoAddSeries_n n).ub k=y, 221 | Real.sqrtTwoAddSeries 0 n=z at * 222 | 223 | have h₈ : (↑y - ↑x) / √(2 - z) ≤ 18 * ↑n / 2 ^ k * 2 ^ n := by 224 | rw [div_eq_mul_inv] 225 | apply mul_le_mul 226 | · rify at h₆ 227 | linarith 228 | · rw [← inv_inv (2^n), 229 | inv_le_inv₀ (by (have : 0 < 2 - z := by linarith); positivity) (by positivity), 230 | ← one_div] 231 | exact h₇ 232 | · positivity 233 | · positivity 234 | 235 | have h₉ : 2 * √(2 - z) / 2 ^ k ≤ 4 / 2 ^ k := by 236 | apply div_le_div₀ zero_le_four ?_ (by positivity) le_rfl 237 | suffices √(2 - z) ≤ 2 by linarith 238 | rw [Real.sqrt_le_left zero_le_two] 239 | linarith 240 | 241 | have h₁₀ : ((18 * n * 2 ^ n + 4) / 2 ^ k : ℝ) = 18 * ↑n / 2 ^ k * 2 ^ n + 4 / 2 ^ k := by 242 | ring_nf 243 | 244 | linarith 245 | 246 | theorem sqrtTwoSubSqrtTwoAddSeries_ub (n k : ℕ) (hk : 3 ≤ k) : 247 | (sqrtTwoSubSqrtTwoAddSeries_n n).ub k - (18 * n * 2 ^ n + 14) / 2 ^ k ≤ √(2 - Real.sqrtTwoAddSeries 0 n) := by 248 | dsimp [ub] 249 | rw [sqrtTwoSubSqrtTwoAddSeries_eq] 250 | 251 | have h₁ := Sqrt.sqrtq_sub_le_sqrt' (2 - Real.sqrtTwoAddSeries 0 n) ((2 - sqrtTwoAddSeries_n n).lub k) k 252 | ⟨?_, ?_⟩ ?_ hk; rotate_left 253 | · change ((2 + -_ : ℚ) : ℝ) ≤ _ 254 | rw [Rat.cast_add, Rat.cast_ofNat, Rat.cast_neg] 255 | apply tsub_le_tsub_left 256 | rw [← ge_iff_le] 257 | convert ComputableℝSeq.hub _ _ 258 | symm 259 | exact IsComputable.prop 260 | · change _ ≤ ((2 + -_ : ℚ) : ℝ) 261 | rw [Rat.cast_add, Rat.cast_ofNat, Rat.cast_neg] 262 | apply tsub_le_tsub_left 263 | convert ComputableℝSeq.hlb _ _ 264 | symm 265 | exact IsComputable.prop 266 | · linarith [Real.sqrtTwoAddSeries_lt_two n] 267 | 268 | have hx : ((2 - sqrtTwoAddSeries_n n).lub k).fst = 2 - (sqrtTwoAddSeries_n n).ub k := by 269 | change (2 +- _) = _ 270 | rw [sub_eq_add_neg] 271 | rfl 272 | have hy : ((2 - sqrtTwoAddSeries_n n).lub k).snd = 2 - (sqrtTwoAddSeries_n n).lb k := by 273 | change (2 +- _) = _ 274 | rw [sub_eq_add_neg] 275 | rfl 276 | simp only [hx, hy, Rat.cast_sub, Rat.cast_ofNat, sub_sub_sub_cancel_left] at h₁ 277 | clear hx hy 278 | have h₄ := Real.sqrtTwoAddSeries_zero_nonneg n 279 | have h₅ := Real.sqrtTwoAddSeries_lt_two n 280 | have h₆ := sqrtTwoAddSeries_n_bounds n k hk 281 | have h₇ := (real_sqrtTwoAddSeries_lb n).le 282 | 283 | generalize 284 | (Sqrt.sqrtq ((2 - sqrtTwoAddSeries_n n).lub k) k).fst=w, 285 | (sqrtTwoAddSeries_n n).lb k=x, 286 | (sqrtTwoAddSeries_n n).ub k=y, 287 | Real.sqrtTwoAddSeries 0 n=z at * 288 | 289 | have h₈ : (↑y - ↑x) / √(2 - z) ≤ 18 * ↑n / 2 ^ k * 2 ^ n := by 290 | rw [div_eq_mul_inv] 291 | apply mul_le_mul 292 | · rify at h₆ 293 | linarith 294 | · rw [← inv_inv (2^n), 295 | inv_le_inv₀ (by (have : 0 < 2 - z := by linarith); positivity) (by positivity), 296 | ← one_div] 297 | exact h₇ 298 | · positivity 299 | · positivity 300 | 301 | have h₉ : 7 * √(2 - z) / 2 ^ k ≤ 14 / 2 ^ k := by 302 | apply div_le_div₀ (by norm_num) ?_ (by positivity) le_rfl 303 | suffices √(2 - z) ≤ 2 by linarith 304 | rw [Real.sqrt_le_left zero_le_two] 305 | linarith 306 | 307 | have h₁₀ : ((18 * n * 2 ^ n + 14) / 2 ^ k : ℝ) = 18 * ↑n / 2 ^ k * 2 ^ n + 14 / 2 ^ k := by 308 | ring_nf 309 | 310 | linarith 311 | 312 | /-- See theorem Real.pi_lt_sqrtTwoAddSeries in Mathlib -/ 313 | def pi_lb (n : ℕ) : ℚ := 314 | 2 ^ (n + 1) * (sqrtTwoSubSqrtTwoAddSeries_n n).lb (3 * n) 315 | 316 | /-- See theorem Real.pi_gt_sqrtTwoAddSeries in Mathlib -/ 317 | def pi_ub (n : ℕ) : ℚ := 318 | 2 ^ (n + 1) * (sqrtTwoSubSqrtTwoAddSeries_n n).ub (3 * n) + 1 / 4 ^ n 319 | 320 | -- ~70ms for 10^-40 precision 321 | -- #time #eval! Rat.toDecimal (prec := 40) (pi_ub 65 - 3.14159265358979323846264338327950288419716939937510) 322 | -- #time #eval! Rat.toDecimal (prec := 40) (pi_lb 65 - 3.14159265358979323846264338327950288419716939937510) 323 | 324 | theorem pi_lb_le_pi (n : ℕ) : pi_lb n ≤ Real.pi := by 325 | refine le_trans ?_ (Real.pi_gt_sqrtTwoAddSeries n).le 326 | simp only [pi_lb, Rat.cast_mul, Rat.cast_pow, Rat.cast_ofNat, Nat.ofNat_pos, pow_pos, mul_le_mul_left] 327 | convert ComputableℝSeq.hlb _ _ 328 | symm 329 | exact IsComputable.prop 330 | 331 | theorem pi_ub_ge_pi (n : ℕ) : Real.pi ≤ pi_ub n := by 332 | refine le_trans (Real.pi_lt_sqrtTwoAddSeries n).le ?_ 333 | simp only [one_div, pi_ub, Rat.cast_add, Rat.cast_mul, Rat.cast_pow, Rat.cast_ofNat, Rat.cast_inv, 334 | add_le_add_iff_right, Nat.ofNat_pos, pow_pos, mul_le_mul_left] 335 | rw [← ge_iff_le] 336 | convert ComputableℝSeq.hub _ _ 337 | symm 338 | exact IsComputable.prop 339 | 340 | theorem pi_lb_ge_pi_sub_pow (n : ℕ) (hn : 0 < n) : Real.pi - 41 * n / 2 ^ n ≤ pi_lb n := by 341 | suffices 2 ^ (n + 1) * ((18 * n * 2 ^ n + 4) / 2 ^ (3 * n)) + 1 / 4 ^ n ≤ (41 * n / 2 ^ n : ℚ) by 342 | rify at this 343 | rw [pi_lb] 344 | have h₁ := sqrtTwoSubSqrtTwoAddSeries_lb n (3 * n) (by omega) 345 | replace h₁ := mul_le_mul_of_nonneg_left h₁ (show 0 ≤ 2 ^ (n + 1) by positivity) 346 | rw [mul_add] at h₁ 347 | have h₂ := Real.pi_lt_sqrtTwoAddSeries n 348 | push_cast 349 | linarith 350 | 351 | by_cases hn' : n < 2 352 | · interval_cases n 353 | norm_num 354 | clear hn 355 | push_neg at hn' 356 | qify at hn' 357 | 358 | have h₁ : (2 ^ (n + 1) * ((18 * n * 2 ^ n + 4) / 2 ^ (3 * n)) + 1 / 4 ^ n : ℚ) 359 | = 36 * n / 2 ^ n + 9 / 4 ^ n := by 360 | rw [show (4 : ℚ) ^ n = (2 ^ 2) ^ n by rfl, ← pow_mul] 361 | field_simp 362 | ring_nf 363 | rw [h₁]; clear h₁ 364 | 365 | suffices (9 / 4 ^ n : ℚ) ≤ 5 * n / 2 ^ n by 366 | simp only [← mul_div] at this ⊢ 367 | linarith 368 | 369 | exact div_le_div₀ (by positivity) (by linarith) (by positivity) (pow_le_pow_left₀ rfl rfl n) 370 | 371 | theorem pi_ub_le_pi_add_pow (n : ℕ) (hn : 0 < n) : pi_ub n ≤ Real.pi + 51 * n / 2 ^ n := by 372 | suffices 2 ^ (n + 1) * ((18 * n * 2 ^ n + 14) / 2 ^ (3 * n)) + 1 / 4 ^ n ≤ (51 * n / 2 ^ n : ℚ) by 373 | rify at this 374 | rw [pi_ub] 375 | have h₁ := sqrtTwoSubSqrtTwoAddSeries_ub n (3 * n) (by omega) 376 | replace h₁ := mul_le_mul_of_nonneg_left h₁ (show 0 ≤ 2 ^ (n + 1) by positivity) 377 | rw [mul_sub] at h₁ 378 | have h₂ := Real.pi_gt_sqrtTwoAddSeries n 379 | push_cast 380 | linarith 381 | 382 | by_cases hn' : n < 2 383 | · interval_cases n 384 | norm_num 385 | clear hn 386 | push_neg at hn' 387 | qify at hn' 388 | 389 | have h₁ : (2 ^ (n + 1) * ((18 * n * 2 ^ n + 14) / 2 ^ (3 * n)) + 1 / 4 ^ n : ℚ) 390 | = 36 * n / 2 ^ n + 29 / 4 ^ n := by 391 | rw [show (4 : ℚ) ^ n = (2 ^ 2) ^ n by rfl, ← pow_mul] 392 | field_simp 393 | ring_nf 394 | rw [h₁]; clear h₁ 395 | 396 | suffices (29 / 4 ^ n : ℚ) ≤ 15 * n / 2 ^ n by 397 | simp only [← mul_div] at this ⊢ 398 | linarith 399 | 400 | exact div_le_div₀ (by positivity) (by linarith) (by positivity) (pow_le_pow_left₀ rfl rfl n) 401 | 402 | theorem pi_lb_causeq : ∃ (h' : IsCauSeq abs pi_lb), Real.mk ⟨pi_lb, h'⟩ = Real.pi := by 403 | refine Real.of_near pi_lb Real.pi ?_ 404 | intro ε hε 405 | have h₁ := Filter.Tendsto.const_mul 41 (tendsto_pow_const_div_const_pow_of_one_lt 1 one_lt_two) 406 | simp only [pow_one, mul_zero] at h₁ 407 | replace h₁ := h₁.eventually_mem (Ioo_mem_nhds (neg_neg_iff_pos.mpr hε) hε) 408 | simp only [Set.mem_Ioo, Filter.eventually_atTop, ge_iff_le] at h₁ 409 | obtain ⟨i, hi⟩ := h₁ 410 | use max i 1 411 | intro j hj 412 | specialize hi j (le_of_max_le_left hj) 413 | rw [abs_lt] 414 | have h₂ := pi_lb_ge_pi_sub_pow j (le_of_max_le_right hj) 415 | rw [← mul_div] at h₂ 416 | constructor 417 | · linarith 418 | · linarith [pi_lb_le_pi j] 419 | 420 | theorem pi_ub_causeq : ∃ (h' : IsCauSeq abs pi_ub), Real.mk ⟨pi_ub, h'⟩ = Real.pi := by 421 | refine Real.of_near pi_ub Real.pi ?_ 422 | intro ε hε 423 | have h₁ := Filter.Tendsto.const_mul 51 (tendsto_pow_const_div_const_pow_of_one_lt 1 one_lt_two) 424 | simp only [pow_one, mul_zero] at h₁ 425 | replace h₁ := h₁.eventually_mem (Ioo_mem_nhds (neg_neg_iff_pos.mpr hε) hε) 426 | simp only [Set.mem_Ioo, Filter.eventually_atTop, ge_iff_le] at h₁ 427 | obtain ⟨i, hi⟩ := h₁ 428 | use max i 1 429 | intro j hj 430 | specialize hi j (le_of_max_le_left hj) 431 | rw [abs_lt] 432 | have h₂ := pi_ub_le_pi_add_pow j (le_of_max_le_right hj) 433 | rw [← mul_div] at h₂ 434 | constructor 435 | · linarith [pi_ub_ge_pi j] 436 | · linarith 437 | 438 | def Pi : ComputableℝSeq := 439 | mk Real.pi 440 | (lub := fun n ↦ ⟨⟨pi_lb n, pi_ub n⟩, 441 | Rat.cast_le.mp <| (pi_lb_le_pi n).trans (pi_ub_ge_pi n)⟩) 442 | (hlb := pi_lb_le_pi) 443 | (hub := pi_ub_ge_pi) 444 | (hcl := pi_lb_causeq.rec (fun w _ ↦ w)) 445 | (hcu := pi_ub_causeq.rec (fun w _ ↦ w)) 446 | (heq := by 447 | obtain ⟨_, h₁⟩ := pi_lb_causeq 448 | obtain ⟨_, h₂⟩ := pi_ub_causeq 449 | rw [← Real.mk_eq, h₁, h₂] 450 | ) 451 | 452 | end Pi 453 | 454 | end ComputableℝSeq 455 | 456 | namespace IsComputable 457 | 458 | instance instComputablePi : IsComputable (Real.pi) where 459 | seq := ComputableℝSeq.Pi 460 | prop := ComputableℝSeq.mk_val_eq_val 461 | 462 | end IsComputable 463 | 464 | example : 465 | 2 < √(Real.pi + 1) 466 | ∧ √(1 - Real.pi) = 0 467 | ∧ 5 * Real.pi / √(2 + Real.pi) < 7 468 | ∧ (31415926 / 10000000 : ℚ) < Real.pi 469 | ∧ Real.pi < 31415927 / 10000000 470 | := by 471 | native_decide 472 | -------------------------------------------------------------------------------- /ComputableReal/SpecialFunctions/Sqrt.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal.IsComputable 2 | import Mathlib.Data.Real.Sqrt 3 | import Mathlib.Analysis.SpecialFunctions.Log.Base 4 | import Mathlib.Data.Real.GoldenRatio 5 | 6 | namespace ComputableℝSeq 7 | 8 | open scoped QInterval 9 | 10 | namespace Sqrt 11 | 12 | theorem boundedSqrt_le_rsqrt (y : ℚ) (n : ℕ) (b : ℕ) (hb : 0 < b): 13 | mkRat (Int.sqrt (y.num * b^n)) ((y.den * b^n).sqrt + 1) ≤ Real.sqrt y := by 14 | simp only [Rat.mkRat_eq_div, Nat.cast_add, Nat.cast_one, Int.cast_add, Int.cast_one] 15 | rify 16 | by_cases hy : y ≤ 0 17 | · have h₁ : √↑y = 0 := by 18 | rw [Real.sqrt_eq_zero'] 19 | exact Rat.cast_nonpos.mpr hy 20 | have h₂ : Int.sqrt (y.num * b ^ n) = 0 := by 21 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 22 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr hy) (by positivity) 23 | simp [h₁, h₂] 24 | push_neg at hy 25 | rw [Rat.cast_def, Real.sqrt_div' _ (Nat.cast_nonneg' y.den)] 26 | conv_rhs => 27 | equals √(↑(y.num * b^n)) / √(↑(y.den * b^n)) => 28 | field_simp 29 | ring_nf 30 | apply div_le_div₀ 31 | · exact Real.sqrt_nonneg _ 32 | · convert Real.nat_sqrt_le_real_sqrt 33 | have h₄ : 0 < y.num * b ^ n := by positivity 34 | have := Int.toNat_of_nonneg h₄.le 35 | rify at this 36 | rw [this] 37 | norm_cast 38 | · simp [← Nat.ne_zero_iff_zero_lt, Nat.sqrt_eq_zero] 39 | positivity 40 | · exact Real.real_sqrt_le_nat_sqrt_succ 41 | 42 | theorem rsqrt_le_boundedSqrt (y : ℚ) (n : ℕ) (b : ℕ) (hb : 0 < b): 43 | Real.sqrt y ≤ mkRat (Int.sqrt (y.num * b^n) + 1) (y.den * b^n).sqrt := by 44 | simp only [Rat.mkRat_eq_div, Nat.cast_add, Nat.cast_one, Int.cast_add, Int.cast_one] 45 | rify 46 | by_cases hy : y ≤ 0 47 | · have h₁ : √↑y = 0 := by 48 | rw [Real.sqrt_eq_zero'] 49 | exact Rat.cast_nonpos.mpr hy 50 | have h₂ : Int.sqrt (y.num * b ^ n) = 0 := by 51 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 52 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr hy) (by positivity) 53 | simp [h₁, h₂] 54 | push_neg at hy 55 | rw [Rat.cast_def, Real.sqrt_div' _ (Nat.cast_nonneg' y.den)] 56 | conv_lhs => 57 | equals √(↑(y.num * b^n)) / √(↑(y.den * b^n)) => 58 | field_simp 59 | ring_nf 60 | apply div_le_div₀ 61 | · have h₁ : 0 < (y.num * b ^ n).sqrt := by 62 | suffices 0 < Nat.sqrt (Int.toNat (y.num) * b ^ n) by 63 | rw [Int.sqrt.eq_1] 64 | norm_cast 65 | convert this 66 | conv_rhs => apply (Int.toNat_ofNat _).symm 67 | push_cast 68 | congr 69 | exact (Int.toNat_of_nonneg ((Rat.num_pos.mpr hy).le)).symm 70 | by_contra h₁ 71 | simp [← Nat.ne_zero_iff_zero_lt, Nat.sqrt_eq_zero, hy, hb.ne'] at h₁ 72 | positivity 73 | · convert Real.real_sqrt_le_nat_sqrt_succ 74 | have h₄ : 0 < y.num * b ^ n := by positivity 75 | have := Int.toNat_of_nonneg h₄.le 76 | rify at this 77 | rw [this] 78 | norm_cast 79 | · simp [← Nat.ne_zero_iff_zero_lt, Nat.sqrt_eq_zero, hb.ne'] 80 | · exact Real.nat_sqrt_le_real_sqrt 81 | 82 | /-- A version of square roots for rational intervals, that give an interval containing the actual 83 | square root, that are at most b^-n wider than the true interval. -/ 84 | def boundedSqrt (x : ℚInterval) (n : ℕ) (b : ℕ) (hb : 0 < b) : ℚInterval := 85 | ⟨⟨ 86 | let q := x.fst; 87 | mkRat (Int.sqrt (q.num * b^n)) ((q.den * b^n).sqrt + 1), 88 | let q := x.snd; 89 | mkRat (Int.sqrt (q.num * b^n) + 1) (q.den * b^n).sqrt⟩, 90 | by 91 | dsimp 92 | rify 93 | trans Real.sqrt (x.snd) 94 | · trans Real.sqrt (x.fst) 95 | · apply boundedSqrt_le_rsqrt _ _ _ hb 96 | · apply Real.sqrt_le_sqrt 97 | exact_mod_cast x.2 98 | · apply rsqrt_le_boundedSqrt _ _ _ hb 99 | ⟩ 100 | 101 | def sqrtq (x : ℚInterval) (n : ℕ) : ℚInterval := 102 | --shortcut with an if to slightly speed things up 103 | if x.snd ≤ 0 then 0 else boundedSqrt x n 4 (by norm_num) 104 | 105 | theorem sqrt_lb_def (q : ℚInterval) (n : ℕ) : 106 | (sqrtq q n).fst = if q.snd ≤ 0 then 0 else 107 | mkRat (Int.sqrt (q.fst.num * 4^n)) ((q.fst.den * 4^n).sqrt + 1) := by 108 | convert apply_ite (fun (x : ℚInterval) ↦ x.fst) _ _ _ 109 | 110 | theorem sqrt_ub_def (q : ℚInterval) (n : ℕ) : 111 | (sqrtq q n).snd = if q.snd ≤ 0 then 0 else 112 | mkRat (Int.sqrt (q.snd.num * 4 ^ n) + 1) (q.snd.den * 4 ^ n).sqrt := by 113 | convert apply_ite (fun (x : ℚInterval) ↦ x.snd) _ _ _ 114 | 115 | theorem sqrtq_nonneg (q : ℚInterval) (n : ℕ) : 0 ≤ (sqrtq q n).fst := by 116 | rw [sqrt_lb_def] 117 | split_ifs 118 | · rfl 119 | · rw [Rat.mkRat_eq_div] 120 | have := Int.sqrt_nonneg (q.toProd.1.num * 4 ^ n) 121 | positivity 122 | 123 | theorem lb_le_sqrt_lb (q : ℚInterval) (n : ℕ) : (sqrtq q n).fst ≤ Real.sqrt q.fst := by 124 | rw [sqrtq] 125 | split_ifs with h 126 | · simp 127 | · apply boundedSqrt_le_rsqrt q.toProd.1 128 | norm_num 129 | 130 | theorem sqrt_ub_le_ub (q : ℚInterval) (n : ℕ) : Real.sqrt q.snd ≤ (sqrtq q n).snd := by 131 | rw [sqrtq] 132 | split_ifs with h 133 | · have := Real.sqrt_eq_zero'.mpr (Rat.cast_nonpos.mpr h) 134 | simp [this] 135 | · apply rsqrt_le_boundedSqrt q.toProd.2 136 | norm_num 137 | 138 | /-- This equality is a generally true way to "split a denominator", but is particularly useful 139 | as an approximation when ε is small compared to y, and we wish to approximate 140 | `x / (y + ε)` with `x / y`. -/ 141 | lemma denom_err (x y ε : ℝ) (hy : y ≠ 0) (hyε : y + ε ≠ 0 ) : 142 | x / (y + ε) = x / y - (x / y) * ε / (y + ε) := by 143 | field_simp 144 | ring_nf 145 | 146 | theorem sqrt_le_mkRat_add (q : ℚ) (n : ℕ) : 147 | Real.sqrt q ≤ mkRat (Int.sqrt (q.num * 4^n)) ((q.den * 4^n).sqrt + 1) + 2 * Real.sqrt q / 2^n := by 148 | nth_rewrite 4 [← Rat.mkRat_self q] 149 | nth_rewrite 1 [← Rat.mkRat_self q] 150 | simp only [Rat.mkRat_eq_div, Rat.cast_div, Rat.cast_intCast, Rat.cast_natCast, Nat.cast_nonneg, 151 | Real.sqrt_div', Nat.cast_add, Nat.cast_one, Rat.cast_add, Rat.cast_one, one_div] 152 | have hd := Rat.den_pos q 153 | generalize q.num = x, q.den = y at * 154 | clear q 155 | rcases le_or_lt x 0 with h|h 156 | · have h₁ : √↑x = 0 := by 157 | rw [Real.sqrt_eq_zero'] 158 | exact Int.cast_nonpos.mpr h 159 | have h₂ : Int.sqrt (x * 4 ^ n) = 0 := by 160 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 161 | exact Int.mul_nonpos_of_nonpos_of_nonneg h (by positivity) 162 | simp [h₁, h₂] 163 | · obtain ⟨z,hz⟩ := Int.eq_ofNat_of_zero_le h.le 164 | subst x 165 | conv_rhs => 166 | enter [1,1,1,1] 167 | tactic => norm_cast 168 | rw [Int.sqrt_natCast] 169 | simp only [Int.cast_natCast] 170 | /- 171 | x/(y+ε) = (x/y) - (x/y - x/(y+ε)) = (x/y) - x*(1/y - 1/(y+ε)) = x/y - x*((y+ε) - y)/(y*(y+ε)) 172 | = x/y - (x/y)*(ε/(y+ε)). The error is thus at most (x/y)*ε/(y+ε), which is upper bounded by 173 | ≤ (sqrt q) * 1 / 4^n. 174 | Similarly for the numerator. 175 | -/ 176 | have h₁ := @Real.floor_real_sqrt_eq_nat_sqrt (z * 4^n) 177 | rify at h₁ 178 | rw [← h₁, ← Int.self_sub_fract] 179 | clear h₁ 180 | have h₂ := Int.fract_lt_one √(↑z * 4 ^ n) 181 | have h₃ := Int.fract_nonneg √(↑z * 4 ^ n) 182 | generalize Int.fract √(↑z * 4 ^ n) = ε₁ at * 183 | 184 | have h₁ := @Real.floor_real_sqrt_eq_nat_sqrt (y * 4^n) 185 | rify at h₁ 186 | rw [← h₁, ← Int.self_sub_fract] 187 | clear h₁ 188 | have h₄ := Int.fract_lt_one √(↑y * 4 ^ n) 189 | have h₅ := Int.fract_nonneg √(↑y * 4 ^ n) 190 | rw [sub_add_comm] 191 | rw [← sub_sub_cancel 1 (Int.fract _)] at h₄ h₅ 192 | generalize 1 - Int.fract √(↑y * 4 ^ n) = ε₂ at * 193 | simp only [sub_lt_self_iff, sub_nonneg] at h₄ h₅ 194 | 195 | rw [Real.sqrt_mul', Real.sqrt_mul', 196 | show (4 ^ n = ((2 ^ n) ^ 2 : ℝ)) by norm_cast; rw [Nat.pow_right_comm]] 197 | rotate_left; positivity; positivity 198 | simp only [Nat.ofNat_nonneg, pow_nonneg, Real.sqrt_sq] 199 | 200 | rw [_root_.add_comm ε₂, sub_div, denom_err] 201 | rotate_left; positivity; positivity 202 | 203 | rw [show √↑z * 2 ^ n / (√↑y * 2 ^ n) = √↑z / √↑y by field_simp; ring_nf] 204 | suffices (√↑z / √↑y * ε₂ / (√↑y * 2 ^ n + ε₂) ≤ √↑z / √↑y / 2 ^ n) 205 | ∧ (ε₁ / (√↑y * 2 ^ n + ε₂) ≤ √↑z / √↑y / 2 ^ n) by 206 | rcases this 207 | rw [← mul_div 2] 208 | linarith 209 | 210 | replace h : 1 ≤ √↑z := Real.one_le_sqrt.mpr (by norm_cast at h ⊢) 211 | replace hd : 1 ≤ √↑y := Real.one_le_sqrt.mpr (Nat.one_le_cast.mpr hd) 212 | 213 | constructor 214 | · apply div_le_div₀ 215 | · positivity 216 | · exact mul_le_of_le_one_right (by positivity) h₅ 217 | · positivity 218 | · trans √↑y * 2 ^ n 219 | · exact le_mul_of_one_le_left (by positivity) hd 220 | · exact le_add_of_nonneg_right h₄.le 221 | · rw [div_div] 222 | apply div_le_div₀ 223 | · positivity 224 | · exact h₂.le.trans h 225 | · positivity 226 | · exact le_add_of_nonneg_right h₄.le 227 | 228 | theorem mkRat_sub_le_sqrt (q : ℚ) (n : ℕ) : 229 | (if q ≤ 0 then 0 else 230 | mkRat (Int.sqrt (q.num * 4 ^ n) + 1) (q.den * 4 ^ n).sqrt 231 | ) - 7 * Real.sqrt q / 2^n ≤ Real.sqrt q := by 232 | split_ifs with h 233 | · rify at h 234 | simp [Real.sqrt_eq_zero'.mpr h] 235 | 236 | push_neg at h 237 | replace h : 0 < q.num := Rat.num_pos.mpr h 238 | nth_rewrite 4 [← Rat.mkRat_self q] 239 | nth_rewrite 3 [← Rat.mkRat_self q] 240 | simp only [Rat.mkRat_eq_div, Rat.cast_div, Rat.cast_intCast, Rat.cast_natCast, Nat.cast_nonneg, 241 | Real.sqrt_div', Nat.cast_add, Nat.cast_one, Rat.cast_add, Rat.cast_one, one_div] 242 | have hd := Rat.den_pos q 243 | generalize q.num = x, q.den = y at * 244 | clear q 245 | 246 | obtain ⟨z,hz⟩ := Int.eq_ofNat_of_zero_le h.le 247 | subst x 248 | 249 | conv_lhs => 250 | enter [1,1,1,1] 251 | tactic => norm_cast 252 | simp only [Int.cast_add, Int.cast_natCast, Int.cast_one] 253 | 254 | replace h : 1 ≤ √↑z := Real.one_le_sqrt.mpr (by norm_cast at h ⊢) 255 | have h2pow : (1 : ℝ) ≤ 2 ^ n := by exact_mod_cast Nat.one_le_two_pow 256 | 257 | have h₁ := @Real.floor_real_sqrt_eq_nat_sqrt (z * 4^n) 258 | rify at h₁ 259 | rw [← h₁, ← Int.self_sub_fract] 260 | clear h₁ 261 | have h₄ := Int.fract_lt_one √(↑z * 4 ^ n) 262 | have h₅ := Int.fract_nonneg √(↑z * 4 ^ n) 263 | rw [sub_add_comm] 264 | rw [← sub_sub_cancel 1 (Int.fract _)] at h₄ h₅ 265 | generalize 1 - Int.fract √(↑z * 4 ^ n) = ε₂ at * 266 | simp only [sub_lt_self_iff, sub_nonneg] at h₄ h₅ 267 | 268 | --Have to special-case the y=1 case. Otherwise the denominator 1 / (√y * 2^n - ε₁) looks like it 269 | -- "could be" arbitrarily close to zero, and so cause a big blowup in error. 270 | by_cases hd' : y = 1 271 | · subst y 272 | simp only [Int.cast_add, Int.cast_one, one_mul, Nat.cast_one, Real.sqrt_one, div_one, 273 | tsub_le_iff_right, ge_iff_le] 274 | rw [show (4 ^ n = ((2 ^ n) ^ 2 : ℕ)) by rw [Nat.pow_right_comm], Nat.sqrt_eq'] 275 | rw [Real.sqrt_mul', show (4 ^ n = ((2 ^ n) ^ 2 : ℝ)) by norm_cast; rw [Nat.pow_right_comm], Real.sqrt_sq, 276 | Nat.cast_pow, Nat.cast_ofNat, add_div] 277 | rotate_left; positivity; positivity 278 | simp only [isUnit_iff_ne_zero, ne_eq, pow_eq_zero_iff', OfNat.ofNat_ne_zero, false_and, 279 | not_false_eq_true, IsUnit.mul_div_cancel_right, _root_.add_comm ( _ / _ ), add_le_add_iff_left] 280 | exact div_le_div_of_nonneg_right (by linarith) (by positivity) 281 | 282 | replace hd' : 2 ≤ y := by omega 283 | replace hd : 1 ≤ √↑y := Real.one_le_sqrt.mpr (Nat.one_le_cast.mpr hd) 284 | 285 | have h₁ := @Real.floor_real_sqrt_eq_nat_sqrt (y * 4^n) 286 | rify at h₁ 287 | rw [← h₁, ← Int.self_sub_fract] 288 | clear h₁ 289 | have h₂ := Int.fract_lt_one √(↑y * 4 ^ n) 290 | have h₃ := Int.fract_nonneg √(↑y * 4 ^ n) 291 | generalize Int.fract √(↑y * 4 ^ n) = ε₁ at * 292 | 293 | rw [Real.sqrt_mul', Real.sqrt_mul', 294 | show (4 ^ n = ((2 ^ n) ^ 2 : ℝ)) by norm_cast; rw [Nat.pow_right_comm]] 295 | rotate_left; positivity; positivity 296 | simp only [Nat.ofNat_nonneg, pow_nonneg, Real.sqrt_sq] 297 | 298 | rw [_root_.add_comm ε₂, add_div, sub_eq_add_neg _ ε₁, denom_err] 299 | rotate_left 300 | · positivity 301 | · nlinarith 302 | 303 | rw [show √↑z * 2 ^ n / (√↑y * 2 ^ n) = √↑z / √↑y by field_simp; ring_nf] 304 | simp only [_root_.mul_neg, neg_div, sub_neg_eq_add] 305 | suffices (√↑z / √↑y * ε₁ / (√↑y * 2 ^ n + -ε₁) ≤ 3 * (√↑z / √↑y / 2 ^ n)) 306 | ∧ (ε₂ / (√↑y * 2 ^ n + -ε₁) ≤ 4 * (√↑z / √↑y / 2 ^ n)) by 307 | rcases this 308 | rw [← mul_div 7] 309 | linarith 310 | 311 | have hi₁ : 1 /3 ≤ √↑y - ε₁ := by 312 | suffices 4 / 3 ≤ √↑y by linarith 313 | trans √2 314 | · rw [Real.le_sqrt' (by positivity)] 315 | norm_num 316 | · exact Real.sqrt_le_sqrt (Nat.ofNat_le_cast.mpr hd') 317 | have hi₂ : 0 < √↑y * 2 ^ n - ε₁ := by 318 | apply lt_of_lt_of_le (show 0 < (1 / 3 : ℝ) by norm_num) 319 | apply hi₁.trans <| sub_le_sub_right (le_mul_of_one_le_right (by positivity) h2pow) _ 320 | 321 | constructor 322 | · ring_nf 323 | rw [mul_assoc, mul_assoc _ _ 3, mul_le_mul_iff_of_pos_left (by positivity)] 324 | apply mul_le_of_le_one_of_le' h₂.le ?_ (by positivity) (by positivity) 325 | field_simp 326 | rw [div_le_div_iff₀ hi₂ (by positivity)] 327 | calc (_ : ℝ) ≤ 3 * (1/3 * (2^n)) := by ring_nf; rfl 328 | _ ≤ 3 * ((√↑y - ε₁) * 2 ^ n) := 329 | mul_le_mul_of_nonneg_left (mul_le_mul_of_nonneg_right hi₁ (by positivity)) zero_le_three 330 | _ = 3 * (√↑y * 2 ^ n - ε₁ * 2 ^ n) := by ring_nf 331 | _ ≤ 3 * (√↑y * 2 ^ n - ε₁) := 332 | mul_le_mul_of_nonneg_left (tsub_le_tsub_left (le_mul_of_one_le_right h₃ h2pow) _) (by positivity) 333 | · rw [div_div, mul_div, ← sub_eq_add_neg] 334 | rw [div_le_div_iff₀ hi₂ (by positivity)] 335 | apply mul_le_of_le_one_of_le' h₅ ?_ (by positivity) (by positivity) 336 | conv_rhs => 337 | equals (√↑z * √↑y * 2 ^ n) + √↑z * (3 * √↑y * 2 ^ n - 4 * ε₁) => 338 | ring_nf 339 | suffices 0 ≤ 3 * √↑y * 2 ^ n - 4 * ε₁ by nlinarith 340 | have : √↑y ≤ √↑y * 2 ^ n := le_mul_of_one_le_right (by positivity) h2pow 341 | linarith 342 | 343 | /-- The square root lower bound has an error that can be composed into one part `2 * √r / 2^n` 344 | that decays with n, but depends on the value; and one part `(x.snd - x.fst) / √r`, that 345 | is proportional to width of the input interval. -/ 346 | theorem sqrt_le_sqrtq_add (r : ℝ) (x : ℚInterval) (n : ℕ) (hq : x.fst ≤ r ∧ r ≤ x.snd) (hx : 0 < x.fst) : 347 | √r ≤ (sqrtq x n).fst + 2 * √r / 2^n + (x.snd - x.fst) / (2 * √x.fst) := by 348 | have hxp : (0 : ℝ) < x.fst := by exact Rat.cast_pos.mpr hx 349 | have hx₂ : 0 < x.snd := by rify; linarith 350 | have hrp : 0 < r := by linarith 351 | have hx21 := sub_nonneg_of_le <| hq.1.trans hq.2 352 | have hrx1 := sub_nonneg_of_le hq.1 353 | suffices (√r - √x.fst) * (1 - 2 / 2^n) ≤ (x.snd - x.fst) / (2 * √x.fst) by 354 | have h₁ := sqrt_le_mkRat_add x.fst n 355 | have h₂ := sqrt_lb_def x n 356 | rw [if_neg (Rat.not_le.mpr hx₂)] at h₂ 357 | rw [← h₂] at h₁ 358 | ring_nf at * 359 | linarith 360 | by_cases hn : n = 0 ∨ n = 1 361 | · rcases hn with rfl|rfl 362 | · norm_num1 363 | trans 0 364 | · simpa using Real.sqrt_le_sqrt hq.1 365 | · positivity 366 | · norm_num 367 | positivity 368 | replace hn : 2 ≤ n := by omega 369 | have h₃ : √r - √x.fst = (r - x.fst) / (√r + √x.fst) := by 370 | field_simp 371 | ring_nf 372 | rw [Real.sq_sqrt hrp.le, Real.sq_sqrt hxp.le] 373 | have h₄ : (0 : ℝ) ≤ 1 - 2 / 2 ^ n := by 374 | have : 2 ≤ 2^n := hn.trans Nat.lt_two_pow_self.le 375 | rify at this 376 | have : (0 : ℝ) ≤ 2^n - 2 := sub_nonneg_of_le this 377 | field_simp 378 | positivity 379 | trans (x.snd - x.fst) / (2 * √x.fst) * (1 - 2 / 2^n) 380 | · rw [h₃] 381 | refine mul_le_mul_of_nonneg_right ?_ h₄ 382 | apply div_le_div₀ hx21 383 | · linarith 384 | · positivity 385 | · linarith [Real.sqrt_le_sqrt hq.1] 386 | · apply mul_le_of_le_one_right 387 | · positivity 388 | · rw [tsub_le_iff_right, le_add_iff_nonneg_right] 389 | positivity 390 | 391 | /-- Similar to `sqrt_le_sqrtq_add`, but doesn't require `0 < x.fst`. -/ 392 | theorem sqrt_le_sqrtq_add' (r : ℝ) (x : ℚInterval) (n : ℕ) (hq : x.fst ≤ r ∧ r ≤ x.snd) (hr : 0 < r): 393 | √r ≤ (sqrtq x n).fst + 2 * √r / 2^n + (x.snd - x.fst) / √r := by 394 | have hx₂ : 0 < x.snd := by rify; linarith 395 | have hx21 := sub_nonneg_of_le <| hq.1.trans hq.2 396 | have hrx1 := sub_nonneg_of_le hq.1 397 | suffices (√r - √x.fst) * (1 - 2 / 2^n) ≤ (x.snd - x.fst) / √r by 398 | have h₁ := sqrt_le_mkRat_add x.fst n 399 | have h₂ := sqrt_lb_def x n 400 | rw [if_neg (Rat.not_le.mpr hx₂)] at h₂ 401 | rw [← h₂] at h₁ 402 | ring_nf at * 403 | linarith 404 | by_cases hn : n = 0 ∨ n = 1 405 | · rcases hn with rfl|rfl 406 | · norm_num1 407 | trans 0 408 | · simpa using Real.sqrt_le_sqrt hq.1 409 | · positivity 410 | · norm_num 411 | positivity 412 | replace hn : 2 ≤ n := by omega 413 | have h₃ : √r - √x.fst ≤ (r - x.fst) / (√r + √x.fst) := by 414 | by_cases hxp : 0 ≤ x.fst 415 | · apply le_of_eq 416 | field_simp 417 | ring_nf 418 | rw [Real.sq_sqrt hr.le, Real.sq_sqrt (Rat.cast_nonneg.mpr hxp)] 419 | · replace hxp : x.fst < (0 : ℝ) := Rat.cast_lt_zero.mpr (Rat.not_le.mp hxp) 420 | simp only [Real.sqrt_eq_zero_of_nonpos (hxp.le), sub_zero, add_zero] 421 | rw [le_div_iff₀ (by positivity), Real.mul_self_sqrt hr.le] 422 | linarith 423 | have h₄ : (0 : ℝ) ≤ 1 - 2 / 2 ^ n := by 424 | have : 2 ≤ 2^n := hn.trans Nat.lt_two_pow_self.le 425 | rify at this 426 | have : (0 : ℝ) ≤ 2^n - 2 := sub_nonneg_of_le this 427 | field_simp 428 | positivity 429 | trans (x.snd - x.fst) / (√r) * (1 - 2 / 2^n) 430 | · refine mul_le_mul_of_nonneg_right ?_ h₄ 431 | apply h₃.trans 432 | apply div_le_div₀ hx21 433 | · linarith 434 | · positivity 435 | · linarith [Real.sqrt_nonneg x.fst] 436 | · apply mul_le_of_le_one_right 437 | · positivity 438 | · rw [tsub_le_iff_right, le_add_iff_nonneg_right] 439 | positivity 440 | 441 | theorem sqrtq_sub_le_sqrt (r : ℝ) (x : ℚInterval) (n : ℕ) (hq : x.fst ≤ r ∧ r ≤ x.snd) (hx : 0 < x.fst) 442 | (hn : 3 ≤ n) : (sqrtq x n).snd - (7 * √r / 2^n) - (x.snd - x.fst) / √x.fst ≤ √r := by 443 | have hxp : (0 : ℝ) < x.fst := by exact Rat.cast_pos.mpr hx 444 | have hx₂ : 0 < x.snd := by rify; linarith 445 | have hxp₂ : (0 : ℝ) < x.snd := by linarith 446 | have hrp : 0 < r := by linarith 447 | have hx21 := sub_nonneg_of_le <| hq.1.trans hq.2 448 | have hrx1 := sub_nonneg_of_le hq.1 449 | suffices (√x.snd - √r) * (1 + 7 / 2^n) ≤ (x.snd - x.fst) / √x.fst by 450 | have h₁ := mkRat_sub_le_sqrt x.snd n 451 | have h₂ := sqrt_ub_def x n 452 | rw [if_neg (Rat.not_le.mpr hx₂)] at h₂ h₁ 453 | rw [← h₂] at h₁ 454 | ring_nf at * 455 | linarith 456 | have h₃ : √x.snd - √r = (x.snd - r) / (√x.snd + √r) := by 457 | field_simp 458 | ring_nf 459 | rw [Real.sq_sqrt hrp.le, Real.sq_sqrt hxp₂.le] 460 | have h₄ : (0 : ℝ) ≤ 1 + 7 / 2 ^ n := by positivity 461 | trans (x.snd - x.fst) / √x.fst * ((1 + 7 / 2^n) / 2) 462 | · suffices (√↑x.toProd.2 - √r) * (1 + 7 / 2 ^ n) ≤ (↑x.toProd.2 - ↑x.toProd.1) / (2 * √↑x.toProd.1) * (1 + 7 / 2 ^ n) by 463 | convert this using 1 464 | ring_nf 465 | rw [h₃] 466 | refine mul_le_mul_of_nonneg_right ?_ h₄ 467 | apply div_le_div₀ hx21 468 | · linarith 469 | · positivity 470 | · linarith [Real.sqrt_le_sqrt hq.1, Real.sqrt_le_sqrt hq.2] 471 | · apply mul_le_of_le_one_right 472 | · positivity 473 | · have : (7 / 2^n : ℝ) ≤ 1 := by 474 | rw [div_le_one₀ (by norm_num)] 475 | apply le_trans (b := 8) (by norm_num) 476 | have := Real.monotone_rpow_of_base_ge_one (one_le_two) (Nat.ofNat_le_cast.mpr hn) 477 | norm_num at this 478 | exact this 479 | linarith 480 | 481 | theorem sqrtq_sub_le_sqrt' (r : ℝ) (x : ℚInterval) (n : ℕ) (hq : x.fst ≤ r ∧ r ≤ x.snd) (hr : 0 < r) 482 | (hn : 3 ≤ n) : (sqrtq x n).snd - (7 * √r / 2^n) - (x.snd - x.fst) / √r ≤ √r := by 483 | have hx₂ : 0 < x.snd := by rify; linarith 484 | have hxp₂ : (0 : ℝ) < x.snd := by linarith 485 | have hx21 := sub_nonneg_of_le <| hq.1.trans hq.2 486 | have hrx1 := sub_nonneg_of_le hq.1 487 | suffices (√x.snd - √r) * (1 + 7 / 2^n) ≤ (x.snd - x.fst) / √r by 488 | have h₁ := mkRat_sub_le_sqrt x.snd n 489 | have h₂ := sqrt_ub_def x n 490 | rw [if_neg (Rat.not_le.mpr hx₂)] at h₂ h₁ 491 | rw [← h₂] at h₁ 492 | ring_nf at * 493 | linarith 494 | have h₃ : √x.snd - √r = (x.snd - r) / (√x.snd + √r) := by 495 | field_simp 496 | ring_nf 497 | rw [Real.sq_sqrt hr.le, Real.sq_sqrt hxp₂.le] 498 | have h₄ : (0 : ℝ) ≤ 1 + 7 / 2 ^ n := by positivity 499 | trans (x.snd - x.fst) / √r * ((1 + 7 / 2^n) / 2) 500 | · suffices (√↑x.toProd.2 - √r) * (1 + 7 / 2 ^ n) ≤ (↑x.toProd.2 - ↑x.toProd.1) / (2 * √r) * (1 + 7 / 2 ^ n) by 501 | convert this using 1 502 | ring_nf 503 | rw [h₃] 504 | refine mul_le_mul_of_nonneg_right ?_ h₄ 505 | apply div_le_div₀ hx21 506 | · linarith 507 | · positivity 508 | · linarith [Real.sqrt_le_sqrt hq.1, Real.sqrt_le_sqrt hq.2] 509 | · apply mul_le_of_le_one_right 510 | · positivity 511 | · have : (7 / 2^n : ℝ) ≤ 1 := by 512 | rw [div_le_one₀ (by norm_num)] 513 | apply le_trans (b := 8) (by norm_num) 514 | have := Real.monotone_rpow_of_base_ge_one (one_le_two) (Nat.ofNat_le_cast.mpr hn) 515 | norm_num at this 516 | exact this 517 | linarith 518 | 519 | theorem TLUW_lower : TendstoLocallyUniformlyWithout 520 | (fun n (x : ℚ) => ↑((fun n q => 521 | mkRat (Int.sqrt (q.num * 4 ^ n)) ((q.den * 4 ^ n).sqrt + 1)) n x)) 522 | (fun q => √↑q) := by 523 | rw [TendstoLocallyUniformlyWithout] 524 | intro ε hε x 525 | dsimp 526 | rcases lt_or_le x 0 with h|h 527 | · use Set.Iic 0, Iic_mem_nhds h, 0 528 | intro b _ y hy 529 | change y ≤ (0:ℝ) at hy 530 | have h₂ : Int.sqrt (y.num * 4 ^ b) = 0 := by 531 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 532 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr <| Rat.cast_nonpos.mp hy) (by positivity) 533 | simp [Real.sqrt_eq_zero'.mpr hy, h₂, hε] 534 | · set tm := max (2 * x) 1 535 | have htm₀ : 0 < tm := by positivity 536 | have htm : x < tm := by 537 | by_cases 0 < x 538 | · exact lt_sup_of_lt_left (by linarith) 539 | · exact lt_sup_of_lt_right (by linarith) 540 | use Set.Ioo (-1) tm, Ioo_mem_nhds (by linarith) htm 541 | set ε' := (ε / (2 * tm.sqrt)) with hε' 542 | set a := Int.clog 2 (1 / ε') with ha 543 | use a.toNat 544 | rintro b hb q ⟨hq₁, hq₂⟩ 545 | by_cases hq₃ : q ≤ 0 546 | · have h₂ : Int.sqrt (q.num * 4 ^ b) = 0 := by 547 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 548 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr hq₃) (by positivity) 549 | simp [Real.sqrt_eq_zero'.mpr (Rat.cast_nonpos.mpr hq₃), h₂, hε] 550 | push_neg at hq₃ 551 | suffices 2 * √↑q / 2 ^ b < ε by 552 | have hb₁ := boundedSqrt_le_rsqrt q b 4 (by norm_num) 553 | rw [Nat.cast_ofNat] at hb₁ 554 | have hb₂ := sqrt_le_mkRat_add q b 555 | rw [abs_sub_lt_iff] 556 | constructor <;> linarith 557 | replace hb : Int.clog 2 (1 / ε') ≤ b := Int.toNat_le.mp hb 558 | replace hb : 2 ^ (Int.clog 2 (1 / ε')) ≤ (2 : ℝ) ^ (b : ℤ) := zpow_le_zpow_right₀ (one_le_two) hb 559 | replace hb := le_trans (Int.self_le_zpow_clog Nat.one_lt_two (1 / ε')) hb 560 | rw [hε', zpow_natCast] at hb 561 | have hqtm := Real.sqrt_lt_sqrt (Rat.cast_nonneg.mpr hq₃.le) hq₂ 562 | have hq0 := Real.sqrt_pos_of_pos (Rat.cast_pos.mpr hq₃) 563 | simp only [one_div, inv_div] at hb 564 | rw [div_le_iff₀ (by positivity)] at hb 565 | rw [div_lt_iff₀ (by positivity), _root_.mul_comm ε] 566 | linarith 567 | 568 | theorem TLUW_upper : TendstoLocallyUniformlyWithout 569 | (fun n (x : ℚ) => ↑((fun n q => 570 | if q ≤ 0 then 0 else mkRat (Int.sqrt (q.num * 4 ^ n) + 1) (q.den * 4 ^ n).sqrt) n x)) 571 | (fun q => √↑q) := by 572 | rw [TendstoLocallyUniformlyWithout] 573 | intro ε hε x 574 | dsimp 575 | rcases lt_or_le x 0 with h|h 576 | · use Set.Iic 0, Iic_mem_nhds h, 0 577 | intro b _ y hy 578 | change y ≤ (0:ℝ) at hy 579 | have hy' := Rat.cast_nonpos.mp hy 580 | have h₂ : Int.sqrt (y.num * 4 ^ b) = 0 := by 581 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 582 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr <| hy') (by positivity) 583 | simp [Real.sqrt_eq_zero'.mpr hy, h₂, hε, hy'] 584 | · set tm := max (2 * x) 1 585 | have htm₀ : 0 < tm := by positivity 586 | have htm : x < tm := by 587 | by_cases 0 < x 588 | · exact lt_sup_of_lt_left (by linarith) 589 | · exact lt_sup_of_lt_right (by linarith) 590 | use Set.Ioo (-1) tm, Ioo_mem_nhds (by linarith) htm 591 | set ε' := (ε / (7 * tm.sqrt)) with hε' 592 | set a := Int.clog 2 (1 / ε') with ha 593 | use a.toNat 594 | rintro b hb q ⟨hq₁, hq₂⟩ 595 | by_cases hq₃ : q ≤ 0 596 | · have h₂ : Int.sqrt (q.num * 4 ^ b) = 0 := by 597 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 598 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr hq₃) (by positivity) 599 | simp [Real.sqrt_eq_zero'.mpr (Rat.cast_nonpos.mpr hq₃), h₂, hε, hq₃] 600 | have hb₂ := mkRat_sub_le_sqrt q b 601 | rw [if_neg hq₃] at hb₂ ⊢ 602 | push_neg at hq₃ 603 | suffices 7 * √↑q / 2 ^ b < ε by 604 | have hb₁ := rsqrt_le_boundedSqrt q b 4 (by norm_num) 605 | rw [Nat.cast_ofNat] at hb₁ 606 | rw [abs_sub_lt_iff] 607 | constructor <;> linarith 608 | replace hb : Int.clog 2 (1 / ε') ≤ b := Int.toNat_le.mp hb 609 | replace hb : 2 ^ (Int.clog 2 (1 / ε')) ≤ (2 : ℝ) ^ (b : ℤ) := zpow_le_zpow_right₀ (one_le_two) hb 610 | replace hb := le_trans (Int.self_le_zpow_clog Nat.one_lt_two (1 / ε')) hb 611 | rw [hε', zpow_natCast] at hb 612 | have hqtm := Real.sqrt_lt_sqrt (Rat.cast_nonneg.mpr hq₃.le) hq₂ 613 | have hq0 := Real.sqrt_pos_of_pos (Rat.cast_pos.mpr hq₃) 614 | simp only [one_div, inv_div] at hb 615 | rw [div_le_iff₀ (by positivity)] at hb 616 | rw [div_lt_iff₀ (by positivity), _root_.mul_comm ε] 617 | linarith 618 | 619 | def sqrt : ComputableℝSeq → ComputableℝSeq := 620 | of_TendstoLocallyUniformly_Continuous 621 | (f := Real.sqrt) 622 | (hf := Real.continuous_sqrt) 623 | (fImpl := fun n q ↦ sqrtq q n) 624 | (fImpl_l := fun n q ↦ mkRat (Int.sqrt (q.num * 4^n)) ((q.den * 4^n).sqrt + 1)) 625 | (fImpl_u := fun n q ↦ if q ≤ 0 then 0 else mkRat (Int.sqrt (q.num * 4^n) + 1) (q.den * 4^n).sqrt) 626 | (hImplDef := by 627 | rintro n ⟨⟨q₁, q₂⟩, hq⟩ 628 | dsimp [sqrtq] 629 | split_ifs 630 | · have h : Int.sqrt (q₁.num * 4 ^ n) = 0 := by 631 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 632 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr (by linarith)) (by positivity) 633 | simp [h]; rfl 634 | · rfl 635 | ) 636 | (hTLU_l := TLUW_lower) (hTLU_u := TLUW_upper) 637 | (hlb := by 638 | intro n ⟨⟨q₁, q₂⟩, hq⟩ x ⟨hx₁, hx₂⟩ 639 | have := lb_le_sqrt_lb ⟨⟨q₁, q₂⟩, hq⟩ n 640 | rw [sqrtq, boundedSqrt] at this 641 | split_ifs at this with h 642 | · have h0 : Int.sqrt (q₁.num * 4 ^ n) = 0 := by 643 | rw [Int.sqrt.eq_1, Int.ofNat_eq_zero, Nat.sqrt_eq_zero, Int.toNat_eq_zero] 644 | exact Int.mul_nonpos_of_nonpos_of_nonneg (Rat.num_nonpos.mpr (by linarith)) (by positivity) 645 | simp [h0] 646 | · exact le_trans this (Real.sqrt_le_sqrt hx₁) 647 | ) 648 | (hub := by 649 | intro n ⟨⟨q₁, q₂⟩, hq⟩ x ⟨hx₁, hx₂⟩ 650 | dsimp at * 651 | split_ifs with h 652 | · suffices √x = 0 by 653 | simp [h, this] 654 | rw [Real.sqrt_eq_zero'] 655 | exact le_trans hx₂ (Rat.cast_nonpos.mpr h) 656 | · have := sqrt_ub_le_ub ⟨⟨q₁, q₂⟩, hq⟩ n 657 | rw [sqrtq, boundedSqrt, if_neg h] at this 658 | exact le_trans (Real.sqrt_le_sqrt hx₂) this 659 | ) 660 | 661 | end Sqrt 662 | 663 | end ComputableℝSeq 664 | 665 | namespace IsComputable 666 | 667 | instance instComputableSqrt (x : ℝ) [hx : IsComputable x] : IsComputable (x.sqrt) := 668 | .lift (Real.sqrt) ComputableℝSeq.Sqrt.sqrt 669 | (by apply ComputableℝSeq.val_of_TendstoLocallyUniformly_Continuous) hx 670 | 671 | instance instComputableGoldenRatio : IsComputable goldenRatio := 672 | inferInstanceAs (IsComputable ((1 + √5) / 2)) 673 | 674 | instance instComputableGoldenConj : IsComputable goldenConj := 675 | inferInstanceAs (IsComputable ((1 - √5) / 2)) 676 | 677 | end IsComputable 678 | -------------------------------------------------------------------------------- /ComputableReal/aux_lemmas.lean: -------------------------------------------------------------------------------- 1 | import Mathlib.Data.Real.Archimedean 2 | 3 | --============ 4 | --silly lemmas 5 | theorem abs_ite_le [inst : LinearOrderedAddCommGroup α] (x : α) : 6 | abs x = if 0 ≤ x then x else -x := by 7 | split_ifs <;> simp_all 8 | next h => 9 | exact LT.lt.le h 10 | 11 | namespace CauSeq 12 | 13 | variable [LinearOrderedField α] {a b : CauSeq α abs} 14 | 15 | theorem sup_equiv_of_equivs (ha : a ≈ c) (hb : b ≈ c) : a ⊔ b ≈ c := by 16 | intro n hn 17 | obtain ⟨i₁, hi₁⟩ := ha n hn 18 | obtain ⟨i₂, hi₂⟩ := hb n hn 19 | use max i₁ i₂ 20 | intro j hj 21 | replace hi₁ := hi₁ j (Nat.max_le.mp hj).1 22 | replace hi₂ := hi₂ j (Nat.max_le.mp hj).2 23 | dsimp at hi₁ hi₂ ⊢ 24 | rw [max_def] 25 | rw [abs_ite_le] at hi₁ hi₂ ⊢ 26 | split_ifs at hi₁ hi₂ ⊢ 27 | all_goals linarith 28 | 29 | theorem equiv_sup_of_equivs (ha : c ≈ a) (hb : c ≈ b) : c ≈ a ⊔ b := 30 | Setoid.symm (sup_equiv_of_equivs (Setoid.symm ha) (Setoid.symm hb)) 31 | 32 | theorem inf_equiv_of_equivs (ha : a ≈ c) (hb : b ≈ c) : a ⊓ b ≈ c := by 33 | --if we had a version of `neg_inf` for CauSeq this could be easily 34 | --reduced to `sup_equiv_of_equivs`. 35 | intro n hn 36 | obtain ⟨i₁, hi₁⟩ := ha n hn 37 | obtain ⟨i₂, hi₂⟩ := hb n hn 38 | use max i₁ i₂ 39 | intro j hj 40 | replace hi₁ := hi₁ j (Nat.max_le.mp hj).1 41 | replace hi₂ := hi₂ j (Nat.max_le.mp hj).2 42 | dsimp at hi₁ hi₂ ⊢ 43 | rw [min_def] 44 | rw [abs_ite_le] at hi₁ hi₂ ⊢ 45 | split_ifs at hi₁ hi₂ ⊢ 46 | all_goals linarith 47 | 48 | /-- Dropping the first n terms of a Cauchy sequence to get a new sequence. -/ 49 | def drop (a : CauSeq α abs) (n : ℕ) : CauSeq α abs := 50 | ⟨fun k ↦ a.val (n+k), fun _ hq ↦ Exists.casesOn (cauchy₂ a hq) 51 | fun i hi ↦ ⟨i, 52 | fun _ hj ↦ hi _ (le_add_of_le_right hj) _ (Nat.le_add_left i n)⟩⟩ 53 | 54 | /-- Dropping elements from a Cauchy sequence gives an equivalent one. -/ 55 | theorem drop_equiv_self (a : CauSeq α abs) (n : ℕ) : a.drop n ≈ a := 56 | fun _ hq ↦ Exists.casesOn (cauchy₂ a hq) 57 | fun i hi ↦ ⟨i, fun _ hj ↦ hi _ (le_add_of_le_right hj) _ hj⟩ 58 | 59 | end CauSeq 60 | 61 | namespace Real 62 | 63 | /-- Every real number has some Caucy sequence converging to it. -/ 64 | theorem exists_CauSeq (x : ℝ) : ∃ s, Real.mk s = x := 65 | let ⟨y,hy⟩ := Quot.exists_rep x.cauchy 66 | ⟨y, by rwa [Real.mk, CauSeq.Completion.mk, Quotient.mk'', Real.ofCauchy.injEq]⟩ 67 | 68 | end Real 69 | 70 | theorem cauchy_real_mk (x : CauSeq ℚ abs) : ∀ ε > 0, ∃ i, ∀ j ≥ i, |x j - Real.mk x| < ε := by 71 | intro ε hε 72 | obtain ⟨q, hq, hq'⟩ := exists_rat_btwn hε 73 | obtain ⟨i, hi⟩ := x.2.cauchy₂ (by simpa using hq) 74 | simp_rw [abs_sub_comm] 75 | refine ⟨i, fun j hj ↦ lt_of_le_of_lt (Real.mk_near_of_forall_near ⟨i, fun k hk ↦ ?_⟩) hq'⟩ 76 | exact_mod_cast (hi k hk j hj).le 77 | 78 | --end silly lemmas 79 | --================ 80 | -------------------------------------------------------------------------------- /ComputableReal/examples.lean: -------------------------------------------------------------------------------- 1 | import ComputableReal 2 | import Mathlib.Probability.Distributions.Gamma 3 | 4 | open Real 5 | 6 | --Casts and basic arithmetic, logic 7 | example : 8 | (10 / 9 : ℝ) < (15 / 2 ⊔ 3) ∧ 9 | |1 / 10 - (1 : ℝ)| < 1 ∧ 10 | (if _ : 1 < 2 then 5 / 2 else 6 : ℝ) < 7 ∧ 11 | (if [1,10].length = 2 then 5 else 10 / 7 : ℝ) / 7 < 1 ∧ 12 | 3.5 < (4 : ℝ) 13 | := by 14 | native_decide 15 | 16 | --Checking the sign of a number 17 | example : (2 - 5 / 2 : ℝ).sign + 1 = 0 := by 18 | native_decide 19 | 20 | --Math with pi 21 | example : 22 | 2 < √(π + 1) 23 | ∧ 5 * π / √(2 + π) < 7 24 | ∧ (31415926 / 10000000 : ℚ) < π 25 | ∧ π < 31415927 / 10000000 26 | := by 27 | native_decide 28 | 29 | --We can prove equalities if values "truncate" -- like Real.sqrt does for negative values. 30 | example : √(1 - π) = 0 := by 31 | native_decide 32 | 33 | --Evaluating the first bit of the harmonic series 34 | example : |∑ x ∈ Finset.range 500, 1 / (x : ℝ) - 6.7908234| < 0.0000001 := by 35 | native_decide 36 | 37 | --Some special functions 38 | example : 39 | 1 / (tanh (1/2) - (1/2)) < -26 ∧ 40 | |4.7 - cosh (sinh (cosh 1))| < 0.02 41 | := by 42 | native_decide 43 | 44 | --e^pi - pi: https://xkcd.com/217/ 45 | example : |exp π - π - 19.9990999| < 0.0000001 := by 46 | native_decide 47 | 48 | --Below are some approximate identities from https://xkcd.com/1047/ 49 | example : |√3 - 2 * exp 1 / π| < 0.002 := by 50 | native_decide 51 | 52 | example : |√5 - (2 / exp 1 + 3 / 2)| < 0.001 := by 53 | native_decide 54 | 55 | --We don't get to cheat! If we try to turn up the "precision" so that it becomes 56 | -- false, the tactic fails and informs us: 57 | /-- 58 | error: tactic 'native_decide' evaluated that the proposition 59 | |√5 - (2 / rexp 1 + 3 / 2)| < 1e-4 60 | is false 61 | -/ 62 | #guard_msgs in 63 | example : |√5 - (2 / exp 1 + 3 / 2)| < 0.0001 := by 64 | native_decide 65 | 66 | open goldenRatio in 67 | example : 68 | let proton_electron_mass_ratio := 1836.152673426; 69 | |proton_electron_mass_ratio - (exp 8 - 10) / φ| < 0.0004 := by 70 | native_decide 71 | 72 | example : |√2 - (3/5 + π / (7 - π))| < 0.00001 := by 73 | native_decide 74 | 75 | /- 76 | If we try to use a function that isn't supported, then the error will sometimes tell us the 77 | relevant function, that it's noncomputable. 78 | -/ 79 | /-- 80 | error: failed to compile definition, consider marking it as 'noncomputable' because it 81 | depends on 'ProbabilityTheory.gammaCDFReal', and it does not have executable code 82 | -/ 83 | #guard_msgs in 84 | example : 0 < ProbabilityTheory.gammaCDFReal 1 1 2 := by 85 | native_decide 86 | 87 | /- 88 | Often, though, it will refer to some *other* noncomputable term. For instance, if you have 89 | division of reals anywhere, it might complain that 'Real.instDivInvMonoid' is noncomputable, 90 | even though `ProbabilityTheory.gammaCDFReal` is the actual culprit. 91 | 92 | This happens because it first tries to make a `Decidable` instance using `ComputableReal`, 93 | it fails (because there's no implementation for `ProbabilityTheory.gammaCDFReal`), and then 94 | it falls back to `Real.decidableLT` (which is really just `Classical.propDecidable`). And 95 | then it tries to compile the whole definition, and fails on the first noncomputable term 96 | it hits. 97 | -/ 98 | /-- 99 | error: failed to compile definition, consider marking it as 'noncomputable' because it depends on 100 | 'Real.instDivInvMonoid', and it does not have executable code 101 | -/ 102 | #guard_msgs in 103 | example : 0 < ProbabilityTheory.gammaCDFReal 1 (1 / 2) 2 := by 104 | native_decide 105 | 106 | /- Operations over complex numbers: -/ 107 | example : (1 + Complex.I) * (1 - Complex.I : ℂ) = 2 := by 108 | native_decide 109 | 110 | example : ‖Complex.I‖ ≠ (1 / 2) := by 111 | native_decide 112 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | Apache License 2 | Version 2.0, January 2004 3 | http://www.apache.org/licenses/ 4 | 5 | TERMS AND CONDITIONS FOR USE, REPRODUCTION, AND DISTRIBUTION 6 | 7 | 1. Definitions. 8 | 9 | "License" shall mean the terms and conditions for use, reproduction, 10 | and distribution as defined by Sections 1 through 9 of this document. 11 | 12 | "Licensor" shall mean the copyright owner or entity authorized by 13 | the copyright owner that is granting the License. 14 | 15 | "Legal Entity" shall mean the union of the acting entity and all 16 | other entities that control, are controlled by, or are under common 17 | control with that entity. For the purposes of this definition, 18 | "control" means (i) the power, direct or indirect, to cause the 19 | direction or management of such entity, whether by contract or 20 | otherwise, or (ii) ownership of fifty percent (50%) or more of the 21 | outstanding shares, or (iii) beneficial ownership of such entity. 22 | 23 | "You" (or "Your") shall mean an individual or Legal Entity 24 | exercising permissions granted by this License. 25 | 26 | "Source" form shall mean the preferred form for making modifications, 27 | including but not limited to software source code, documentation 28 | source, and configuration files. 29 | 30 | "Object" form shall mean any form resulting from mechanical 31 | transformation or translation of a Source form, including but 32 | not limited to compiled object code, generated documentation, 33 | and conversions to other media types. 34 | 35 | "Work" shall mean the work of authorship, whether in Source or 36 | Object form, made available under the License, as indicated by a 37 | copyright notice that is included in or attached to the work 38 | (an example is provided in the Appendix below). 39 | 40 | "Derivative Works" shall mean any work, whether in Source or Object 41 | form, that is based on (or derived from) the Work and for which the 42 | editorial revisions, annotations, elaborations, or other modifications 43 | represent, as a whole, an original work of authorship. For the purposes 44 | of this License, Derivative Works shall not include works that remain 45 | separable from, or merely link (or bind by name) to the interfaces of, 46 | the Work and Derivative Works thereof. 47 | 48 | "Contribution" shall mean any work of authorship, including 49 | the original version of the Work and any modifications or additions 50 | to that Work or Derivative Works thereof, that is intentionally 51 | submitted to Licensor for inclusion in the Work by the copyright owner 52 | or by an individual or Legal Entity authorized to submit on behalf of 53 | the copyright owner. For the purposes of this definition, "submitted" 54 | means any form of electronic, verbal, or written communication sent 55 | to the Licensor or its representatives, including but not limited to 56 | communication on electronic mailing lists, source code control systems, 57 | and issue tracking systems that are managed by, or on behalf of, the 58 | Licensor for the purpose of discussing and improving the Work, but 59 | excluding communication that is conspicuously marked or otherwise 60 | designated in writing by the copyright owner as "Not a Contribution." 61 | 62 | "Contributor" shall mean Licensor and any individual or Legal Entity 63 | on behalf of whom a Contribution has been received by Licensor and 64 | subsequently incorporated within the Work. 65 | 66 | 2. Grant of Copyright License. Subject to the terms and conditions of 67 | this License, each Contributor hereby grants to You a perpetual, 68 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 69 | copyright license to reproduce, prepare Derivative Works of, 70 | publicly display, publicly perform, sublicense, and distribute the 71 | Work and such Derivative Works in Source or Object form. 72 | 73 | 3. Grant of Patent License. Subject to the terms and conditions of 74 | this License, each Contributor hereby grants to You a perpetual, 75 | worldwide, non-exclusive, no-charge, royalty-free, irrevocable 76 | (except as stated in this section) patent license to make, have made, 77 | use, offer to sell, sell, import, and otherwise transfer the Work, 78 | where such license applies only to those patent claims licensable 79 | by such Contributor that are necessarily infringed by their 80 | Contribution(s) alone or by combination of their Contribution(s) 81 | with the Work to which such Contribution(s) was submitted. If You 82 | institute patent litigation against any entity (including a 83 | cross-claim or counterclaim in a lawsuit) alleging that the Work 84 | or a Contribution incorporated within the Work constitutes direct 85 | or contributory patent infringement, then any patent licenses 86 | granted to You under this License for that Work shall terminate 87 | as of the date such litigation is filed. 88 | 89 | 4. Redistribution. You may reproduce and distribute copies of the 90 | Work or Derivative Works thereof in any medium, with or without 91 | modifications, and in Source or Object form, provided that You 92 | meet the following conditions: 93 | 94 | (a) You must give any other recipients of the Work or 95 | Derivative Works a copy of this License; and 96 | 97 | (b) You must cause any modified files to carry prominent notices 98 | stating that You changed the files; and 99 | 100 | (c) You must retain, in the Source form of any Derivative Works 101 | that You distribute, all copyright, patent, trademark, and 102 | attribution notices from the Source form of the Work, 103 | excluding those notices that do not pertain to any part of 104 | the Derivative Works; and 105 | 106 | (d) If the Work includes a "NOTICE" text file as part of its 107 | distribution, then any Derivative Works that You distribute must 108 | include a readable copy of the attribution notices contained 109 | within such NOTICE file, excluding those notices that do not 110 | pertain to any part of the Derivative Works, in at least one 111 | of the following places: within a NOTICE text file distributed 112 | as part of the Derivative Works; within the Source form or 113 | documentation, if provided along with the Derivative Works; or, 114 | within a display generated by the Derivative Works, if and 115 | wherever such third-party notices normally appear. The contents 116 | of the NOTICE file are for informational purposes only and 117 | do not modify the License. You may add Your own attribution 118 | notices within Derivative Works that You distribute, alongside 119 | or as an addendum to the NOTICE text from the Work, provided 120 | that such additional attribution notices cannot be construed 121 | as modifying the License. 122 | 123 | You may add Your own copyright statement to Your modifications and 124 | may provide additional or different license terms and conditions 125 | for use, reproduction, or distribution of Your modifications, or 126 | for any such Derivative Works as a whole, provided Your use, 127 | reproduction, and distribution of the Work otherwise complies with 128 | the conditions stated in this License. 129 | 130 | 5. Submission of Contributions. Unless You explicitly state otherwise, 131 | any Contribution intentionally submitted for inclusion in the Work 132 | by You to the Licensor shall be under the terms and conditions of 133 | this License, without any additional terms or conditions. 134 | Notwithstanding the above, nothing herein shall supersede or modify 135 | the terms of any separate license agreement you may have executed 136 | with Licensor regarding such Contributions. 137 | 138 | 6. Trademarks. This License does not grant permission to use the trade 139 | names, trademarks, service marks, or product names of the Licensor, 140 | except as required for reasonable and customary use in describing the 141 | origin of the Work and reproducing the content of the NOTICE file. 142 | 143 | 7. Disclaimer of Warranty. Unless required by applicable law or 144 | agreed to in writing, Licensor provides the Work (and each 145 | Contributor provides its Contributions) on an "AS IS" BASIS, 146 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or 147 | implied, including, without limitation, any warranties or conditions 148 | of TITLE, NON-INFRINGEMENT, MERCHANTABILITY, or FITNESS FOR A 149 | PARTICULAR PURPOSE. You are solely responsible for determining the 150 | appropriateness of using or redistributing the Work and assume any 151 | risks associated with Your exercise of permissions under this License. 152 | 153 | 8. Limitation of Liability. In no event and under no legal theory, 154 | whether in tort (including negligence), contract, or otherwise, 155 | unless required by applicable law (such as deliberate and grossly 156 | negligent acts) or agreed to in writing, shall any Contributor be 157 | liable to You for damages, including any direct, indirect, special, 158 | incidental, or consequential damages of any character arising as a 159 | result of this License or out of the use or inability to use the 160 | Work (including but not limited to damages for loss of goodwill, 161 | work stoppage, computer failure or malfunction, or any and all 162 | other commercial damages or losses), even if such Contributor 163 | has been advised of the possibility of such damages. 164 | 165 | 9. Accepting Warranty or Additional Liability. While redistributing 166 | the Work or Derivative Works thereof, You may choose to offer, 167 | and charge a fee for, acceptance of support, warranty, indemnity, 168 | or other liability obligations and/or rights consistent with this 169 | License. However, in accepting such obligations, You may act only 170 | on Your own behalf and on Your sole responsibility, not on behalf 171 | of any other Contributor, and only if You agree to indemnify, 172 | defend, and hold each Contributor harmless for any liability 173 | incurred by, or claims asserted against, such Contributor by reason 174 | of your accepting any such warranty or additional liability. 175 | 176 | END OF TERMS AND CONDITIONS 177 | 178 | APPENDIX: How to apply the Apache License to your work. 179 | 180 | To apply the Apache License to your work, attach the following 181 | boilerplate notice, with the fields enclosed by brackets "[]" 182 | replaced with your own identifying information. (Don't include 183 | the brackets!) The text should be enclosed in the appropriate 184 | comment syntax for the file format. We also recommend that a 185 | file or class name and description of purpose be included on the 186 | same "printed page" as the copyright notice for easier 187 | identification within third-party archives. 188 | 189 | Copyright [yyyy] [name of copyright owner] 190 | 191 | Licensed under the Apache License, Version 2.0 (the "License"); 192 | you may not use this file except in compliance with the License. 193 | You may obtain a copy of the License at 194 | 195 | http://www.apache.org/licenses/LICENSE-2.0 196 | 197 | Unless required by applicable law or agreed to in writing, software 198 | distributed under the License is distributed on an "AS IS" BASIS, 199 | WITHOUT WARRANTIES OR CONDITIONS OF ANY KIND, either express or implied. 200 | See the License for the specific language governing permissions and 201 | limitations under the License. 202 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | This is a package for [Lean4](https://github.com/leanprover/lean4) to enable computations with real numbers. As a theorem prover, Lean happily will talk about _all_ real numbers, including ones that cannot be computed. But most real numbers of interest can be computed. This is a package to help automatically determine certain statements about those real numbers. As a concrete example, statements like 2 | ```lean 3 | example : Real.sqrt (1 + (1/4 : ℚ)) > Real.sqrt (2 + Real.sqrt (3 : ℕ)) / (Real.sqrt 7 - 1/2) := by 4 | native_decide 5 | 6 | example : |√3 - 2 * exp 1 / π| < 0.002 := by 7 | native_decide 8 | 9 | example : (1 + Complex.I) * (1 - Complex.I : ℂ) = 2 := by 10 | native_decide 11 | ``` 12 | can now be proved by `native_decide`. 13 | 14 | ## Background 15 | 16 | Colloquially, a real number is often called "computable" if there is an algorithm that produces successive lower- and upper-bounds which converge to that number (although precise definitions by experts may vary), and a computable function is one that maps computable numbers to computable numbers. These sequences are all rational numbers, so that we can represent them exactly. This gives it the flavor of [Interval arithmetic](https://en.wikipedia.org/wiki/Interval_arithmetic), but with the additional guarantee that our intervals need to converge. 17 | 18 | For example, real addition is computable, since we can simply add the respective lower and upper bound sequences for the two summands. But computing the sign of a real is, strictly speaking, not computable: the lower bounds `-1 / 2^n` and upper bounds `1 / 2^n` converge to zero, but at any finite iteration we can't rule out that they're converging to a positive or negative value. For this reason, our implementation is a _partial function_: when the sign is needed, it will continue evaluating all the sequence until it determines the sign; this may not terminate. 19 | 20 | ## Implementation details 21 | 22 | A `ComputableℝSeq` carries the functions for lower and upper bounds, together with proofs of their validity: they are equivalent Cauchy sequences, and the lower bound (resp. upper) stays below (resp. above) the completion of the Cauchy sequence. Addition, negation, multiplication, and continuous functions can be computed on these sequences without issue. 23 | 24 | When the sign function is needed, `ComputableℝSeq.sign` gives a `partial def` for computing this. Evaluating `ComputableℝSeq.sign` on `4 * (1/4 : ℚ) - 1` will immediately give `0`, since the lower and upper bounds will both evaluate to zero; but `Real.sqrt 2 - Real.sqrt 2` will never terminate. Computing the inverse of a `ComputableℝSeq` then calls `sign` appropriately. This also lets us create a `Decidable (x ≤ y)` instance where `x` and `y` are `ComputableℝSeq`. 25 | 26 | To make this painlessly usable with `native_decide`, a typeclass `IsComputable (x : ℝ)` is available. This looks at the expression and tries to _structurally_ synthesize a `ComputableℝSeq` equal to the expression. This means that `Real.sqrt (2 + Real.sqrt 3)` is synthesized from a series of `IsComputable` instances for addition and square roots; but a local assumption like `x = 3` cannot be used. Then, given a pair of `IsComputable` instances on `x : ℝ` and `y : ℝ`, we have `Decidable (x ≤ y)`. 27 | 28 | The partial aspect means that, in general, native_decide will be able to prove any true _inequality_ or _nonequality_, but is only able to provide _equalities_ if they're structurally rational expressions. 29 | 30 | ```lean 31 | example : Real.sqrt 2 > 4 / 3 := by --true inequality, good 32 | native_decide 33 | 34 | example : Real.sqrt 2 ≤ 1 + Real.sqrt 5 := by --true inequality, good 35 | native_decide 36 | 37 | example : Real.sqrt 2 ≠ 4 / 3 := by --true nonequality, good 38 | native_decide 39 | 40 | example : Real.sqrt 2 > 1 + Real.sqrt 5 := by --false inequality, gives an error that it evaluates to false 41 | native_decide 42 | 43 | example : 2 + 5 = 14 / 2 := by --true equality, works because only rational numbers are involved 44 | native_decide 45 | 46 | example : Real.sqrt 2 = Real.sqrt 2 := by --hangs, never terminate 47 | native_decide 48 | 49 | example : Real.sqrt 2 - Real.sqrt 2 = 0:= by --hangs 50 | native_decide 51 | ``` 52 | 53 | There are more examples in [examples.lean](./ComputableReal/examples.lean). 54 | 55 | ## What's supported 56 | * Addition, negation, subtraction 57 | * Multiplication, natural number powers 58 | * Inverses, division, integer powers 59 | * Casts from naturals, rationals, `ofNat` literals, and `OfScientific` literals (e.g `1.2`) 60 | * Simple functions: `ite`, `dite`, `Real.sign`, `abs`, `max`, `min`, `Finset.sum` 61 | * Special functions: `Real.sqrt`, `Real.exp`, `Real.sinh`, `Real.cosh`, `Real.tanh` 62 | * Constants: `Real.pi`, e (as `Real.exp 1`), `goldenRatio`, `goldenConj` 63 | * Complex numbers: field operations, casts, `norm`, `inner`, `normSq` 64 | 65 | ## TODO List 66 | Roughly in order: 67 | * Adding support for `Real.log` - will require careful sign handling similarly to inverses, due to the discontinuity at zero. 68 | * This gets us `Real.arsinh` ... not that it's in particularly high demand. 69 | * Also `Real.logb`, `Real.negMulLog`, `Real.posLog`. 70 | * Adding support for `Real.cos` 71 | * This gets us `Real.sin`, `Real.tan`, and most importantly, real powers `_ ^ _`. These rely on cosine because of their behavior with negative powers. 72 | * Then using `Real.exp` and `Real.sin`/`Real.cos`, we get `Complex.exp`; then we get trig functions on the complex numbers for free too. 73 | * Adding `Real.arctan`, possibly via `Real.two_mul_arctan_add_pi`. 74 | * This gets us `Real.arcsin` and `Real.arccos`, but also importantly `Complex.arg` and `Complex.log`. Then from `Complex.log` we get all the trig functions' inverses on the complex numbers. 75 | * Other, low-priority functions that can be implemented: 76 | * `Real.Gamma`, with the theorem `Real.GammaSeq_tendsto_Gamma`. Actually, given an implementation of `Real.pi`, this would give an alternate implementation of `Real.sin` using `Real.Gamma_mul_Gamma_one_sub` ... but that's probably not very practical. 77 | * Integrals with certain properties (such as positivity / convexity) -------------------------------------------------------------------------------- /lake-manifest.json: -------------------------------------------------------------------------------- 1 | {"version": "1.1.0", 2 | "packagesDir": ".lake/packages", 3 | "packages": 4 | [{"url": "https://github.com/leanprover-community/mathlib4.git", 5 | "type": "git", 6 | "subDir": null, 7 | "scope": "", 8 | "rev": "2e4bea3590e14768d39b454d51bc438cf1731503", 9 | "name": "mathlib", 10 | "manifestFile": "lake-manifest.json", 11 | "inputRev": null, 12 | "inherited": false, 13 | "configFile": "lakefile.lean"}, 14 | {"url": "https://github.com/leanprover-community/plausible", 15 | "type": "git", 16 | "subDir": null, 17 | "scope": "leanprover-community", 18 | "rev": "59a8514bb0ee5bae2689d8be717b5272c9b3dc1c", 19 | "name": "plausible", 20 | "manifestFile": "lake-manifest.json", 21 | "inputRev": "main", 22 | "inherited": true, 23 | "configFile": "lakefile.toml"}, 24 | {"url": "https://github.com/leanprover-community/LeanSearchClient", 25 | "type": "git", 26 | "subDir": null, 27 | "scope": "leanprover-community", 28 | "rev": "003ff459cdd85de551f4dcf95cdfeefe10f20531", 29 | "name": "LeanSearchClient", 30 | "manifestFile": "lake-manifest.json", 31 | "inputRev": "main", 32 | "inherited": true, 33 | "configFile": "lakefile.toml"}, 34 | {"url": "https://github.com/leanprover-community/import-graph", 35 | "type": "git", 36 | "subDir": null, 37 | "scope": "leanprover-community", 38 | "rev": "5013810061a18ca1f5510106172b94c6fbd0a2fc", 39 | "name": "importGraph", 40 | "manifestFile": "lake-manifest.json", 41 | "inputRev": "main", 42 | "inherited": true, 43 | "configFile": "lakefile.toml"}, 44 | {"url": "https://github.com/leanprover-community/ProofWidgets4", 45 | "type": "git", 46 | "subDir": null, 47 | "scope": "leanprover-community", 48 | "rev": "8fff3f074da9237cd4e179fd6dd89be6c4022d41", 49 | "name": "proofwidgets", 50 | "manifestFile": "lake-manifest.json", 51 | "inputRev": "v0.0.52-pre", 52 | "inherited": true, 53 | "configFile": "lakefile.lean"}, 54 | {"url": "https://github.com/leanprover-community/aesop", 55 | "type": "git", 56 | "subDir": null, 57 | "scope": "leanprover-community", 58 | "rev": "ba9a63be53f16b3b6e4043641c6bad4883e650b4", 59 | "name": "aesop", 60 | "manifestFile": "lake-manifest.json", 61 | "inputRev": "master", 62 | "inherited": true, 63 | "configFile": "lakefile.toml"}, 64 | {"url": "https://github.com/leanprover-community/quote4", 65 | "type": "git", 66 | "subDir": null, 67 | "scope": "leanprover-community", 68 | "rev": "7b3b0c8327b3c0214ac49ca6d6922edbb81ab8c9", 69 | "name": "Qq", 70 | "manifestFile": "lake-manifest.json", 71 | "inputRev": "master", 72 | "inherited": true, 73 | "configFile": "lakefile.toml"}, 74 | {"url": "https://github.com/leanprover-community/batteries", 75 | "type": "git", 76 | "subDir": null, 77 | "scope": "leanprover-community", 78 | "rev": "24cbb071689802fd6d3ff42198b19b125004c4e3", 79 | "name": "batteries", 80 | "manifestFile": "lake-manifest.json", 81 | "inputRev": "main", 82 | "inherited": true, 83 | "configFile": "lakefile.toml"}, 84 | {"url": "https://github.com/leanprover/lean4-cli", 85 | "type": "git", 86 | "subDir": null, 87 | "scope": "leanprover", 88 | "rev": "a2eb24a3dbf681f2b655f82ba5ee5b139d4a5abc", 89 | "name": "Cli", 90 | "manifestFile": "lake-manifest.json", 91 | "inputRev": "main", 92 | "inherited": true, 93 | "configFile": "lakefile.toml"}], 94 | "name": "computableReal", 95 | "lakeDir": ".lake"} 96 | -------------------------------------------------------------------------------- /lakefile.lean: -------------------------------------------------------------------------------- 1 | import Lake 2 | open Lake DSL 3 | 4 | package «computableReal» { 5 | -- add any package configuration options here 6 | } 7 | 8 | require mathlib from git 9 | "https://github.com/leanprover-community/mathlib4.git" 10 | 11 | @[default_target] 12 | lean_lib «ComputableReal» { 13 | -- add any library configuration options here 14 | } 15 | -------------------------------------------------------------------------------- /lean-toolchain: -------------------------------------------------------------------------------- 1 | leanprover/lean4:v4.17.0-rc1 2 | --------------------------------------------------------------------------------