Exponential Power Series #
This file defines the exponential power series exp A = ∑ Xⁿ/n! over ℚ-algebras and develops
its key properties, including the fundamental differential equation (exp A)' = exp A,
a uniqueness characterization, and the functional equation for multiplication.
Main definitions #
PowerSeries.exp: The exponential power seriesexp A = ∑ Xⁿ/n!over a ℚ-algebraA.
Main results #
PowerSeries.coeff_exp: The coefficient ofexp Aatnis1/n!.PowerSeries.constantCoeff_exp: The constant term ofexp Ais1.PowerSeries.map_exp:expis preserved by ring homomorphisms between ℚ-algebras.PowerSeries.derivative_exp: The derivative of exp equals exp:d⁄dX A (exp A) = exp A.PowerSeries.exp_unique_of_derivative_eq_self: A power series with derivative equal to itself and constant term1must beexp.PowerSeries.isUnit_exp:exp Ais a unit (invertible).PowerSeries.order_exp: The order ofexp Ais0.PowerSeries.exp_mul_exp_eq_exp_add: The functional equatione^{aX} * e^{bX} = e^{(a+b)X}.PowerSeries.exp_mul_exp_neg_eq_one: The identitye^X * e^{-X} = 1.PowerSeries.exp_pow_eq_rescale_exp: Powers of exp satisfy(e^X)^k = e^{kX}.PowerSeries.exp_pow_sum: Formula for the sum of powers ofexp.
Power series for the exponential function at zero.
Equations
- PowerSeries.exp A = PowerSeries.mk fun (n : ℕ) => (algebraMap ℚ A) (1 / ↑n.factorial)
Instances For
@[simp]
Derivative of exp #
Uniqueness characterization #
theorem
PowerSeries.exp_unique_of_derivative_eq_self
{A : Type u_3}
[CommRing A]
[Algebra ℚ A]
[IsAddTorsionFree A]
{f : PowerSeries A}
(hd : (derivative A) f = f)
(hc : constantCoeff f = 1)
:
A power series with derivative equal to itself and constant term 1 must be exp.
The proof uses induction on coefficients: if f' = f and f(0) = 1, then
coeff (n+1) f * (n+1) = coeff n f, which determines all coefficients uniquely.
Order and invertibility #
@[simp]
theorem
PowerSeries.exp_pow_sum
{A : Type u_4}
[CommRing A]
[Algebra ℚ A]
(n : ℕ)
:
∑ k ∈ Finset.range n, exp A ^ k = mk fun (p : ℕ) => ∑ k ∈ Finset.range n, ↑k ^ p * (algebraMap ℚ A) (↑p.factorial)⁻¹
Shows that $\sum_{k = 0}^{n - 1} (e^{X})^k = \sum_{p = 0}^{\infty} \sum_{k = 0}^{n - 1} \frac{k^p}{p!}X^p$.