Skip to content

Commit e1a19e3

Browse files
committed
feat: dot notation for IsTheta.add_isLittleO, and add_commed variants (#10386)
BREAKING CHANGE: Change `IsTheta.add_isLittleO` into a dot-notation statement, in line with the existing `IsBigO.add_isLittleO`. Move the current `IsTheta.add_isLittleO` statement to `IsLittleO.right_isTheta_add'`, in line with the existing `IsLittleO.right_isBigO_add`. feat: Add `add_comm`ed variants of related lemmas. These changes smoothen the flow when proving e.g. `a + b + c + d + e + f =Θ[l] d`. Co-authored-by: L Lllvvuu <git@llllvvuu.dev>
1 parent 8658c57 commit e1a19e3

File tree

2 files changed

+19
-3
lines changed

2 files changed

+19
-3
lines changed

Mathlib/Analysis/Asymptotics/Asymptotics.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2147,6 +2147,10 @@ theorem IsLittleO.right_isBigO_add {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂)
21472147
((h.def' one_half_pos).right_le_add_of_lt_one one_half_lt_one).isBigO
21482148
#align asymptotics.is_o.right_is_O_add Asymptotics.IsLittleO.right_isBigO_add
21492149

2150+
theorem IsLittleO.right_isBigO_add' {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
2151+
f₂ =O[l] (f₂ + f₁) :=
2152+
add_comm f₁ f₂ ▸ h.right_isBigO_add
2153+
21502154
/-- If `f x = O(g x)` along `cofinite`, then there exists a positive constant `C` such that
21512155
`‖f x‖ ≤ C * ‖g x‖` whenever `g x ≠ 0`. -/
21522156
theorem bound_of_isBigO_cofinite (h : f =O[cofinite] g'') :

Mathlib/Analysis/Asymptotics/Theta.lean

Lines changed: 15 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -326,8 +326,20 @@ alias ⟨IsTheta.of_const_mul_right, IsTheta.const_mul_right⟩ := isTheta_const
326326
#align asymptotics.is_Theta.of_const_mul_right Asymptotics.IsTheta.of_const_mul_right
327327
#align asymptotics.is_Theta.const_mul_right Asymptotics.IsTheta.const_mul_right
328328

329-
lemma IsTheta.add_isLittleO {f₁ f₂ : α → E'}
330-
(h : f₂ =o[l] f₁) : (f₁ + f₂) =Θ[l] f₁ :=
331-
⟨(isBigO_refl _ _).add_isLittleO h, by rw [add_comm]; exact h.right_isBigO_add⟩
329+
theorem IsLittleO.right_isTheta_add {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
330+
f₂ =Θ[l] (f₁ + f₂) :=
331+
⟨h.right_isBigO_add, h.add_isBigO (isBigO_refl _ _)⟩
332+
333+
theorem IsLittleO.right_isTheta_add' {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
334+
f₂ =Θ[l] (f₂ + f₁) :=
335+
add_comm f₁ f₂ ▸ h.right_isTheta_add
336+
337+
lemma IsTheta.add_isLittleO {f₁ f₂ : α → E'} {g : α → F}
338+
(hΘ : f₁ =Θ[l] g) (ho : f₂ =o[l] g) : (f₁ + f₂) =Θ[l] g :=
339+
(ho.trans_isTheta hΘ.symm).right_isTheta_add'.symm.trans hΘ
340+
341+
lemma IsLittleO.add_isTheta {f₁ f₂ : α → E'} {g : α → F}
342+
(ho : f₁ =o[l] g) (hΘ : f₂ =Θ[l] g) : (f₁ + f₂) =Θ[l] g :=
343+
add_comm f₁ f₂ ▸ hΘ.add_isLittleO ho
332344

333345
end Asymptotics

0 commit comments

Comments
 (0)