|
| 1 | +/- |
| 2 | +Copyright (c) 2020 Anatole Dedecker. All rights reserved. |
| 3 | +Released under Apache 2.0 license as described in the file LICENSE. |
| 4 | +Authors: Anatole Dedecker |
| 5 | +-/ |
| 6 | +import analysis.asymptotics |
| 7 | +import analysis.normed_space.ordered |
| 8 | +import analysis.normed_space.bounded_linear_maps |
| 9 | + |
| 10 | +/-! |
| 11 | +# Asymptotic equivalence |
| 12 | +
|
| 13 | +In this file, we define the relation `is_equivalent u v l`, which means that `u-v` is little o of |
| 14 | +`v` along the filter `l`. |
| 15 | +
|
| 16 | +Unlike `is_[oO]` relations, this one requires `u` and `v` to have the same codomaine `β`. While the |
| 17 | +definition only requires `β` to be a `normed_group`, most interesting properties require it to be a |
| 18 | +`normed_field`. |
| 19 | +
|
| 20 | +## Notations |
| 21 | +
|
| 22 | +We introduce the notation `u ~[l] v := is_equivalent u v l`, which you can use by opening the |
| 23 | +`asymptotics` locale. |
| 24 | +
|
| 25 | +## Main results |
| 26 | +
|
| 27 | +If `β` is a `normed_group` : |
| 28 | +
|
| 29 | +- `_ ~[l] _` is an equivalence relation |
| 30 | +- Equivalent statements for `u ~[l] const _ c` : |
| 31 | + - If `c ≠ 0`, this is true iff `tendsto u l (𝓝 c)` (see `is_equivalent_const_iff_tendsto`) |
| 32 | + - For `c = 0`, this is true iff `u =ᶠ[l] 0` (see `is_equivalent_zero_iff_eventually_zero`) |
| 33 | +
|
| 34 | +If `β` is a `normed_field` : |
| 35 | +
|
| 36 | +- Alternative characterization of the relation (see `is_equivalent_iff_exists_eq_mul`) : |
| 37 | +
|
| 38 | + `u ~[l] v ↔ ∃ (φ : α → β) (hφ : tendsto φ l (𝓝 1)), u =ᶠ[l] φ * v` |
| 39 | +
|
| 40 | +- Provided some non-vanishing hypothesis, this can be seen as `u ~[l] v ↔ tendsto (u/v) l (𝓝 1)` |
| 41 | + (see `is_equivalent_iff_tendsto_one`) |
| 42 | +- For any constant `c`, `u ~[l] v` implies `tendsto u l (𝓝 c) ↔ tendsto v l (𝓝 c)` |
| 43 | + (see `is_equivalent.tendsto_nhds_iff`) |
| 44 | +- `*` and `/` are compatible with `_ ~[l] _` (see `is_equivalent.mul` and `is_equivalent.div`) |
| 45 | +
|
| 46 | +If `β` is a `normed_linear_ordered_field` : |
| 47 | +
|
| 48 | +- If `u ~[l] v`, we have `tendsto u l at_top ↔ tendsto v l at_top` |
| 49 | + (see `is_equivalent.tendsto_at_top_iff`) |
| 50 | +
|
| 51 | +-/ |
| 52 | + |
| 53 | +namespace asymptotics |
| 54 | + |
| 55 | +open filter function |
| 56 | +open_locale topological_space |
| 57 | + |
| 58 | +section normed_group |
| 59 | + |
| 60 | +variables {α β : Type*} [normed_group β] |
| 61 | + |
| 62 | +/-- Two functions `u` and `v` are said to be asymptotically equivalent along a filter `l` when |
| 63 | + `u x - v x = o(v x)` as x converges along `l`. -/ |
| 64 | +def is_equivalent (u v : α → β) (l : filter α) := is_o (u - v) v l |
| 65 | + |
| 66 | +localized "notation u ` ~[`:50 l:50 `] `:0 v:50 := asymptotics.is_equivalent u v l" in asymptotics |
| 67 | + |
| 68 | +variables {u v w : α → β} {l : filter α} |
| 69 | + |
| 70 | +lemma is_equivalent.is_o (h : u ~[l] v) : is_o (u - v) v l := h |
| 71 | + |
| 72 | +lemma is_equivalent.is_O (h : u ~[l] v) : is_O u v l := |
| 73 | +(is_O.congr_of_sub h.is_O.symm).mp (is_O_refl _ _) |
| 74 | + |
| 75 | +lemma is_equivalent.is_O_symm (h : u ~[l] v) : is_O v u l := |
| 76 | +begin |
| 77 | + convert h.is_o.right_is_O_add, |
| 78 | + ext, |
| 79 | + simp |
| 80 | +end |
| 81 | + |
| 82 | +@[refl] lemma is_equivalent.refl : u ~[l] u := |
| 83 | +begin |
| 84 | + rw [is_equivalent, sub_self], |
| 85 | + exact is_o_zero _ _ |
| 86 | +end |
| 87 | + |
| 88 | +@[symm] lemma is_equivalent.symm (h : u ~[l] v) : v ~[l] u := |
| 89 | +(h.is_o.trans_is_O h.is_O_symm).symm |
| 90 | + |
| 91 | +@[trans] lemma is_equivalent.trans (huv : u ~[l] v) (hvw : v ~[l] w) : u ~[l] w := |
| 92 | +(huv.is_o.trans_is_O hvw.is_O).triangle hvw.is_o |
| 93 | + |
| 94 | +lemma is_equivalent_zero_iff_eventually_zero : u ~[l] 0 ↔ u =ᶠ[l] 0 := |
| 95 | +begin |
| 96 | + rw [is_equivalent, sub_zero], |
| 97 | + exact is_o_zero_right_iff |
| 98 | +end |
| 99 | + |
| 100 | +lemma is_equivalent_const_iff_tendsto {c : β} (h : c ≠ 0) : u ~[l] const _ c ↔ tendsto u l (𝓝 c) := |
| 101 | +begin |
| 102 | + rw [is_equivalent, is_o_const_iff h], |
| 103 | + split; intro h; |
| 104 | + [ { have := h.add tendsto_const_nhds, rw zero_add at this }, |
| 105 | + { have := h.add tendsto_const_nhds, rw ← sub_self c} ]; |
| 106 | + convert this; ext; simp [sub_eq_add_neg] |
| 107 | +end |
| 108 | + |
| 109 | +lemma is_equivalent.tendsto_const {c : β} (hu : u ~[l] const _ c) : tendsto u l (𝓝 c) := |
| 110 | +begin |
| 111 | + rcases (em $ c = 0) with ⟨rfl, h⟩, |
| 112 | + { exact (tendsto_congr' $ is_equivalent_zero_iff_eventually_zero.mp hu).mpr tendsto_const_nhds }, |
| 113 | + { exact (is_equivalent_const_iff_tendsto h).mp hu } |
| 114 | +end |
| 115 | + |
| 116 | +lemma is_equivalent.tendsto_nhds {c : β} (huv : u ~[l] v) (hu : tendsto u l (𝓝 c)) : |
| 117 | + tendsto v l (𝓝 c) := |
| 118 | +begin |
| 119 | + by_cases h : c = 0, |
| 120 | + { rw [h, ← is_o_one_iff ℝ] at *, |
| 121 | + convert (huv.symm.is_o.trans hu).add hu, |
| 122 | + simp }, |
| 123 | + { rw ← is_equivalent_const_iff_tendsto h at hu ⊢, |
| 124 | + exact huv.symm.trans hu } |
| 125 | +end |
| 126 | + |
| 127 | +lemma is_equivalent.tendsto_nhds_iff {c : β} (huv : u ~[l] v) : |
| 128 | + tendsto u l (𝓝 c) ↔ tendsto v l (𝓝 c) := ⟨huv.tendsto_nhds, huv.symm.tendsto_nhds⟩ |
| 129 | + |
| 130 | +end normed_group |
| 131 | + |
| 132 | +open_locale asymptotics |
| 133 | + |
| 134 | +section normed_field |
| 135 | + |
| 136 | +variables {α β : Type*} [normed_field β] {t u v w : α → β} {l : filter α} |
| 137 | + |
| 138 | +lemma is_equivalent_iff_exists_eq_mul : u ~[l] v ↔ |
| 139 | + ∃ (φ : α → β) (hφ : tendsto φ l (𝓝 1)), u =ᶠ[l] φ * v := |
| 140 | +begin |
| 141 | + rw [is_equivalent, is_o_iff_exists_eq_mul], |
| 142 | + split; rintros ⟨φ, hφ, h⟩; [use (φ + 1), use (φ - 1)]; split, |
| 143 | + { conv in (𝓝 _) { rw ← zero_add (1 : β) }, |
| 144 | + exact hφ.add (tendsto_const_nhds) }, |
| 145 | + { convert h.add (eventually_eq.refl l v); ext; simp [add_mul] }, |
| 146 | + { conv in (𝓝 _) { rw ← sub_self (1 : β) }, |
| 147 | + exact hφ.sub (tendsto_const_nhds) }, |
| 148 | + { convert h.sub (eventually_eq.refl l v); ext; simp [sub_mul] } |
| 149 | +end |
| 150 | + |
| 151 | +lemma is_equivalent.exists_eq_mul (huv : u ~[l] v) : |
| 152 | + ∃ (φ : α → β) (hφ : tendsto φ l (𝓝 1)), u =ᶠ[l] φ * v := |
| 153 | +is_equivalent_iff_exists_eq_mul.mp huv |
| 154 | + |
| 155 | +lemma is_equivalent_of_tendsto_one (hz : ∀ᶠ x in l, v x = 0 → u x = 0) |
| 156 | + (huv : tendsto (u/v) l (𝓝 1)) : u ~[l] v := |
| 157 | +begin |
| 158 | + rw is_equivalent_iff_exists_eq_mul, |
| 159 | + refine ⟨u/v, huv, hz.mono $ λ x hz', (div_mul_cancel_of_imp hz').symm⟩, |
| 160 | +end |
| 161 | + |
| 162 | +lemma is_equivalent_of_tendsto_one' (hz : ∀ x, v x = 0 → u x = 0) (huv : tendsto (u/v) l (𝓝 1)) : |
| 163 | + u ~[l] v := |
| 164 | +is_equivalent_of_tendsto_one (eventually_of_forall hz) huv |
| 165 | + |
| 166 | +lemma is_equivalent_iff_tendsto_one (hz : ∀ᶠ x in l, v x ≠ 0) : |
| 167 | + u ~[l] v ↔ tendsto (u/v) l (𝓝 1) := |
| 168 | +begin |
| 169 | + split, |
| 170 | + { intro hequiv, |
| 171 | + have := hequiv.is_o.tendsto_0, |
| 172 | + simp only [pi.sub_apply, sub_div] at this, |
| 173 | + have key : tendsto (λ x, v x / v x) l (𝓝 1), |
| 174 | + { exact (tendsto_congr' $ hz.mono $ λ x hnz, @div_self _ _ (v x) hnz).mpr tendsto_const_nhds }, |
| 175 | + convert this.add key, |
| 176 | + { ext, simp }, |
| 177 | + { norm_num } }, |
| 178 | + { exact is_equivalent_of_tendsto_one (hz.mono $ λ x hnvz hz, (hnvz hz).elim) } |
| 179 | +end |
| 180 | + |
| 181 | +end normed_field |
| 182 | + |
| 183 | +section smul |
| 184 | + |
| 185 | +lemma is_equivalent.smul {α E 𝕜 : Type*} [normed_field 𝕜] [normed_group E] |
| 186 | + [normed_space 𝕜 E] {a b : α → 𝕜} {u v : α → E} {l : filter α} (hab : a ~[l] b) (huv : u ~[l] v) : |
| 187 | + (λ x, a x • u x) ~[l] (λ x, b x • v x) := |
| 188 | +begin |
| 189 | + rcases hab.exists_eq_mul with ⟨φ, hφ, habφ⟩, |
| 190 | + have : (λ (x : α), a x • u x) - (λ (x : α), b x • v x) =ᶠ[l] λ x, b x • ((φ x • u x) - v x), |
| 191 | + { convert (habφ.comp₂ (•) $ eventually_eq.refl _ u).sub (eventually_eq.refl _ (λ x, b x • v x)), |
| 192 | + ext, |
| 193 | + rw [pi.mul_apply, mul_comm, mul_smul, ← smul_sub] }, |
| 194 | + refine (is_o_congr this.symm $ eventually_eq.refl _ _).mp ((is_O_refl b l).smul_is_o _), |
| 195 | + |
| 196 | + rcases huv.is_O.exists_pos with ⟨C, hC, hCuv⟩, |
| 197 | + rw is_equivalent at *, |
| 198 | + rw is_o_iff at *, |
| 199 | + rw is_O_with at hCuv, |
| 200 | + simp only [metric.tendsto_nhds, dist_eq_norm] at hφ, |
| 201 | + intros c hc, |
| 202 | + specialize hφ ((c/2)/C) (div_pos (by linarith) hC), |
| 203 | + specialize huv (show 0 < c/2, by linarith), |
| 204 | + refine hφ.mp (huv.mp $ hCuv.mono $ λ x hCuvx huvx hφx, _), |
| 205 | + |
| 206 | + have key := |
| 207 | + calc ∥φ x - 1∥ * ∥u x∥ |
| 208 | + ≤ (c/2) / C * ∥u x∥ : mul_le_mul_of_nonneg_right hφx.le (norm_nonneg $ u x) |
| 209 | + ... ≤ (c/2) / C * (C*∥v x∥) : mul_le_mul_of_nonneg_left hCuvx (div_pos (by linarith) hC).le |
| 210 | + ... = c/2 * ∥v x∥ : by {field_simp [hC.ne.symm], ring}, |
| 211 | + |
| 212 | + calc ∥((λ (x : α), φ x • u x) - v) x∥ |
| 213 | + = ∥(φ x - 1) • u x + (u x - v x)∥ : by simp [sub_smul, sub_add] |
| 214 | + ... ≤ ∥(φ x - 1) • u x∥ + ∥u x - v x∥ : norm_add_le _ _ |
| 215 | + ... = ∥φ x - 1∥ * ∥u x∥ + ∥u x - v x∥ : by rw norm_smul |
| 216 | + ... ≤ c / 2 * ∥v x∥ + ∥u x - v x∥ : add_le_add_right key _ |
| 217 | + ... ≤ c / 2 * ∥v x∥ + c / 2 * ∥v x∥ : add_le_add_left huvx _ |
| 218 | + ... = c * ∥v x∥ : by ring, |
| 219 | +end |
| 220 | + |
| 221 | +end smul |
| 222 | + |
| 223 | +section mul_inv |
| 224 | + |
| 225 | +variables {α β : Type*} [normed_field β] {t u v w : α → β} {l : filter α} |
| 226 | + |
| 227 | +lemma is_equivalent.mul (htu : t ~[l] u) (hvw : v ~[l] w) : t * v ~[l] u * w := |
| 228 | +htu.smul hvw |
| 229 | + |
| 230 | +lemma is_equivalent.inv (huv : u ~[l] v) : (λ x, (u x)⁻¹) ~[l] (λ x, (v x)⁻¹) := |
| 231 | +begin |
| 232 | + rw is_equivalent_iff_exists_eq_mul at *, |
| 233 | + rcases huv with ⟨φ, hφ, h⟩, |
| 234 | + rw ← inv_one, |
| 235 | + refine ⟨λ x, (φ x)⁻¹, tendsto.inv' hφ (by norm_num) , _⟩, |
| 236 | + convert h.inv, |
| 237 | + ext, |
| 238 | + simp [mul_inv'] |
| 239 | +end |
| 240 | + |
| 241 | +lemma is_equivalent.div (htu : t ~[l] u) (hvw : v ~[l] w) : |
| 242 | + (λ x, t x / v x) ~[l] (λ x, u x / w x) := |
| 243 | +htu.mul hvw.inv |
| 244 | + |
| 245 | +end mul_inv |
| 246 | + |
| 247 | +section normed_linear_ordered_field |
| 248 | + |
| 249 | +variables {α β : Type*} [normed_linear_ordered_field β] {u v : α → β} {l : filter α} |
| 250 | + |
| 251 | +lemma is_equivalent.tendsto_at_top [order_topology β] (huv : u ~[l] v) (hu : tendsto u l at_top) : |
| 252 | + tendsto v l at_top := |
| 253 | +let ⟨φ, hφ, h⟩ := huv.symm.exists_eq_mul in |
| 254 | +tendsto.congr' h.symm ((mul_comm u φ) ▸ (tendsto_mul_at_top zero_lt_one hu hφ)) |
| 255 | + |
| 256 | +lemma is_equivalent.tendsto_at_top_iff [order_topology β] (huv : u ~[l] v) : |
| 257 | + tendsto u l at_top ↔ tendsto v l at_top := ⟨huv.tendsto_at_top, huv.symm.tendsto_at_top⟩ |
| 258 | + |
| 259 | +end normed_linear_ordered_field |
| 260 | + |
| 261 | +end asymptotics |
0 commit comments