Skip to content

Commit

Permalink
feat: asymptotics lemmas (#17394)
Browse files Browse the repository at this point in the history
  • Loading branch information
Ruben-VandeVelde and MichaelStollBayreuth committed Oct 5, 2024
1 parent 875821f commit b5895d8
Showing 1 changed file with 10 additions and 0 deletions.
10 changes: 10 additions & 0 deletions Mathlib/Analysis/Asymptotics/Asymptotics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -1251,6 +1251,9 @@ theorem IsLittleO.trans_tendsto (hfg : f'' =o[l] g'') (hg : Tendsto g'' l (𝓝
Tendsto f'' l (𝓝 0) :=
hfg.isBigO.trans_tendsto hg

lemma isLittleO_id_one [One F''] [NeZero (1 : F'')] : (fun x : E'' => x) =o[𝓝 0] (1 : E'' → F'') :=
isLittleO_id_const one_ne_zero

/-! ### Multiplication by a constant -/


Expand Down Expand Up @@ -1915,6 +1918,13 @@ theorem isBigO_atTop_iff_eventually_exists_pos {α : Type*}
f =O[atTop] g ↔ ∀ᶠ n₀ in atTop, ∃ c > 0, ∀ n ≥ n₀, c * ‖f n‖ ≤ ‖g n‖ := by
simp_rw [isBigO_iff'', ← exists_prop, Subtype.exists', exists_eventually_atTop]

lemma isBigO_mul_iff_isBigO_div {f g h : α → 𝕜} (hf : ∀ᶠ x in l, f x ≠ 0) :
(fun x ↦ f x * g x) =O[l] h ↔ g =O[l] (fun x ↦ h x / f x) := by
rw [isBigO_iff', isBigO_iff']
refine ⟨fun ⟨c, hc, H⟩ ↦ ⟨c, hc, ?_⟩, fun ⟨c, hc, H⟩ ↦ ⟨c, hc, ?_⟩⟩ <;>
· refine H.congr <| Eventually.mp hf <| Eventually.of_forall fun x hx ↦ ?_
rw [norm_mul, norm_div, ← mul_div_assoc, le_div_iff₀' (norm_pos_iff.mpr hx)]

end Asymptotics

open Asymptotics
Expand Down

0 comments on commit b5895d8

Please sign in to comment.