Documentation

Mathlib.Analysis.SpecialFunctions.Complex.LogBounds

Estimates for the complex logarithm #

We show that log (1+z) differs from its Taylor polynomial up to degree n by at most ‖z‖^(n+1)/((n+1)*(1-‖z‖)) when ‖z‖ < 1; see Complex.norm_log_sub_logTaylor_le.

To this end, we derive the representation of log (1+z) as the integral of 1/(1+tz) over the unit interval (Complex.log_eq_integral) and introduce notation Complex.logTaylor n for the Taylor polynomial up to degree n-1.

TODO #

Refactor using general Taylor series theory, once this exists in Mathlib.

Integral representation of the complex log #

theorem Complex.log_eq_integral {z : } (hz : 1 + z Complex.slitPlane) :
Complex.log (1 + z) = z * ∫ (t : ) in 0 ..1, (1 + t z)⁻¹

Represent log (1 + z) as an integral over the unit interval

theorem Complex.log_inv_eq_integral {z : } (hz : 1 - z Complex.slitPlane) :
Complex.log (1 - z)⁻¹ = z * ∫ (t : ) in 0 ..1, (1 - t z)⁻¹

Represent log (1 - z)⁻¹ as an integral over the unit interval

The Taylor polynomials of the logarithm #

noncomputable def Complex.logTaylor (n : ) :

The nth Taylor polynomial of log at 1, as a function ℂ → ℂ

Equations
Instances For
    theorem Complex.logTaylor_succ (n : ) :
    Complex.logTaylor (n + 1) = Complex.logTaylor n + fun (z : ) => (-1) ^ (n + 1) * z ^ n / n
    theorem Complex.hasDerivAt_logTaylor (n : ) (z : ) :
    HasDerivAt (Complex.logTaylor (n + 1)) (Finset.sum (Finset.range n) fun (j : ) => (-1) ^ j * z ^ j) z

    Bounds for the difference between log and its Taylor polynomials #

    theorem Complex.hasDerivAt_log_sub_logTaylor (n : ) {z : } (hz : 1 + z Complex.slitPlane) :
    HasDerivAt (fun (z : ) => Complex.log (1 + z) - Complex.logTaylor (n + 1) z) ((-z) ^ n * (1 + z)⁻¹) z
    theorem Complex.norm_one_add_mul_inv_le {t : } (ht : t Set.Icc 0 1) {z : } (hz : z < 1) :
    (1 + t * z)⁻¹ (1 - z)⁻¹

    Give a bound on ‖(1 + t * z)⁻¹‖ for 0 ≤ t ≤ 1 and ‖z‖ < 1.

    theorem Complex.integrable_pow_mul_norm_one_add_mul_inv (n : ) {z : } (hz : z < 1) :
    IntervalIntegrable (fun (t : ) => t ^ n * (1 + t * z)⁻¹) MeasureTheory.volume 0 1
    theorem Complex.norm_log_sub_logTaylor_le (n : ) {z : } (hz : z < 1) :
    Complex.log (1 + z) - Complex.logTaylor (n + 1) z z ^ (n + 1) * (1 - z)⁻¹ / (n + 1)

    The difference of log (1+z) and its (n+1)st Taylor polynomial can be bounded in terms of ‖z‖.

    The difference log (1+z) - z is bounded by ‖z‖^2/(2*(1-‖z‖)) when ‖z‖ < 1.

    The difference of log (1-z)⁻¹ and its (n+1)st Taylor polynomial can be bounded in terms of ‖z‖.

    The difference log (1-z)⁻¹ - z is bounded by ‖z‖^2/(2*(1-‖z‖)) when ‖z‖ < 1.

    theorem Complex.hasSum_taylorSeries_log {z : } (hz : z < 1) :
    HasSum (fun (n : ) => (-1) ^ (n + 1) * z ^ n / n) (Complex.log (1 + z))

    The Taylor series of the complex logarithm at 1 converges to the logarithm in the open unit disk.

    theorem Complex.hasSum_taylorSeries_neg_log {z : } (hz : z < 1) :
    HasSum (fun (n : ) => z ^ n / n) (-Complex.log (1 - z))

    The series ∑ z^n/n converges to -log (1-z) on the open unit disk.