Skip to content

Commit

Permalink
feat: dot notation for IsTheta.add_isLittleO, and add_commed vari…
Browse files Browse the repository at this point in the history
…ants (#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>
  • Loading branch information
2 people authored and riccardobrasca committed Feb 18, 2024
1 parent 9740c67 commit e958b98
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 3 deletions.
4 changes: 4 additions & 0 deletions Mathlib/Analysis/Asymptotics/Asymptotics.lean
Original file line number Diff line number Diff line change
Expand Up @@ -2147,6 +2147,10 @@ theorem IsLittleO.right_isBigO_add {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂)
((h.def' one_half_pos).right_le_add_of_lt_one one_half_lt_one).isBigO
#align asymptotics.is_o.right_is_O_add Asymptotics.IsLittleO.right_isBigO_add

theorem IsLittleO.right_isBigO_add' {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
f₂ =O[l] (f₂ + f₁) :=
add_comm f₁ f₂ ▸ h.right_isBigO_add

/-- If `f x = O(g x)` along `cofinite`, then there exists a positive constant `C` such that
`‖f x‖ ≤ C * ‖g x‖` whenever `g x ≠ 0`. -/
theorem bound_of_isBigO_cofinite (h : f =O[cofinite] g'') :
Expand Down
18 changes: 15 additions & 3 deletions Mathlib/Analysis/Asymptotics/Theta.lean
Original file line number Diff line number Diff line change
Expand Up @@ -326,8 +326,20 @@ alias ⟨IsTheta.of_const_mul_right, IsTheta.const_mul_right⟩ := isTheta_const
#align asymptotics.is_Theta.of_const_mul_right Asymptotics.IsTheta.of_const_mul_right
#align asymptotics.is_Theta.const_mul_right Asymptotics.IsTheta.const_mul_right

lemma IsTheta.add_isLittleO {f₁ f₂ : α → E'}
(h : f₂ =o[l] f₁) : (f₁ + f₂) =Θ[l] f₁ :=
⟨(isBigO_refl _ _).add_isLittleO h, by rw [add_comm]; exact h.right_isBigO_add⟩
theorem IsLittleO.right_isTheta_add {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
f₂ =Θ[l] (f₁ + f₂) :=
⟨h.right_isBigO_add, h.add_isBigO (isBigO_refl _ _)⟩

theorem IsLittleO.right_isTheta_add' {f₁ f₂ : α → E'} (h : f₁ =o[l] f₂) :
f₂ =Θ[l] (f₂ + f₁) :=
add_comm f₁ f₂ ▸ h.right_isTheta_add

lemma IsTheta.add_isLittleO {f₁ f₂ : α → E'} {g : α → F}
(hΘ : f₁ =Θ[l] g) (ho : f₂ =o[l] g) : (f₁ + f₂) =Θ[l] g :=
(ho.trans_isTheta hΘ.symm).right_isTheta_add'.symm.trans hΘ

lemma IsLittleO.add_isTheta {f₁ f₂ : α → E'} {g : α → F}
(ho : f₁ =o[l] g) (hΘ : f₂ =Θ[l] g) : (f₁ + f₂) =Θ[l] g :=
add_comm f₁ f₂ ▸ hΘ.add_isLittleO ho

end Asymptotics

0 comments on commit e958b98

Please sign in to comment.