The logarithm as a limit of powers #
This file shows that the logarithm can be expressed as a limit of powers, namely that
p⁻¹ * (x ^ p - 1) tends to log x as p tends to zero for positive x.
Main declarations #
tendstoUniformlyOn_rpow_sub_one_log:p⁻¹ * (x ^ p - 1)tends uniformly tolog xon compact subsets ofIoi 0asptends to zerotendsto_rpow_sub_one_log:p⁻¹ * (x ^ p - 1): the analogous statement for pointwise convergence.
theorem
Real.tendstoLocallyUniformlyOn_rpow_sub_one_log :
TendstoLocallyUniformlyOn (fun (p x : ℝ) => p⁻¹ * (x ^ p - 1)) log (nhdsWithin 0 (Set.Ioi 0)) (Set.Ioi 0)
theorem
tendsto_rpow_sub_one_log
{x : ℝ}
(hx : 0 < x)
:
Filter.Tendsto (fun (p : ℝ) => p⁻¹ * (x ^ p - 1)) (nhdsWithin 0 (Set.Ioi 0)) (nhds (Real.log x))