Radius of convergence of a power series #
This file introduces the notion of the radius of convergence of a power series.
Main definitions #
Let p
be a formal multilinear series from E
to F
, i.e., p n
is a multilinear map on E^n
for n : ℕ
.
p.radius
: the largestr : ℝ≥0∞
such that‖p n‖ * r^n
grows subexponentially.p.le_radius_of_bound
,p.le_radius_of_bound_nnreal
,p.le_radius_of_isBigO
: if‖p n‖ * r ^ n
is bounded above, thenr ≤ p.radius
;p.isLittleO_of_lt_radius
,p.norm_mul_pow_le_mul_pow_of_lt_radius
,p.isLittleO_one_of_lt_radius
,p.norm_mul_pow_le_of_lt_radius
,p.nnnorm_mul_pow_le_of_lt_radius
: ifr < p.radius
, then‖p n‖ * r ^ n
tends to zero exponentially;p.lt_radius_of_isBigO
: ifr ≠ 0
and‖p n‖ * r ^ n = O(a ^ n)
for some-1 < a < 1
, thenr < p.radius
;p.partialSum n x
: the sum∑_{i = 0}^{n-1} pᵢ xⁱ
.p.sum x
: the sum∑'_{i = 0}^{∞} pᵢ xⁱ
.
Implementation details #
We only introduce the radius of convergence of a power series, as p.radius
.
For a power series in finitely many dimensions, there is a finer (directional, coordinate-dependent)
notion, describing the polydisk of convergence. This notion is more specific, and not necessary to
build the general theory. We do not define it here.
Given a formal multilinear series p
and a vector x
, then p.sum x
is the sum Σ pₙ xⁿ
. A
priori, it only behaves well when ‖x‖ < p.radius
.
Instances For
Given a formal multilinear series p
and a vector x
, then p.partialSum n x
is the sum
Σ pₖ xᵏ
for k ∈ {0,..., n-1}
.
Equations
- p.partialSum n x = ∑ k ∈ Finset.range n, (p k) fun (x_1 : Fin k) => x
Instances For
The partial sums of a formal multilinear series are continuous.
The radius of a formal multilinear series #
The radius of a formal multilinear series is the largest r
such that the sum Σ ‖pₙ‖ ‖y‖ⁿ
converges for all ‖y‖ < r
. This implies that Σ pₙ yⁿ
converges for all ‖y‖ < r
, but these
definitions are not equivalent in general.
Instances For
If ‖pₙ‖ rⁿ
is bounded in n
, then the radius of p
is at least r
.
If ‖pₙ‖ rⁿ
is bounded in n
, then the radius of p
is at least r
.
If ‖pₙ‖ rⁿ = O(1)
, as n → ∞
, then the radius of p
is at least r
.
0
has infinite radius of convergence
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ
tends to zero exponentially:
for some 0 < a < 1
, ‖p n‖ rⁿ = o(aⁿ)
.
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ = o(1)
.
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ
tends to zero exponentially:
for some 0 < a < 1
and C > 0
, ‖p n‖ * r ^ n ≤ C * a ^ n
.
If r ≠ 0
and ‖pₙ‖ rⁿ = O(aⁿ)
for some -1 < a < 1
, then r < p.radius
.
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ
is bounded.
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ
is bounded.
For r
strictly smaller than the radius of p
, then ‖pₙ‖ rⁿ
is bounded.
If the radius of p
is positive, then ‖pₙ‖
grows at most geometrically.
The radius of the sum of two formal series is at least the minimum of their two radii.
This is a version of radius_compContinuousLinearMap_linearIsometryEquiv_eq
with better
opportunity for unification, at the cost of manually supplying some hypotheses.