Skip to content
This repository was archived by the owner on Jul 24, 2024. It is now read-only.

Commit c70feeb

Browse files
committed
feat(analysis/analytic/inverse): convergence of the inverse of a power series (#5854)
If a formal multilinear series has a positive radius of convergence, then its inverse also does.
1 parent 19f24ce commit c70feeb

File tree

6 files changed

+335
-2
lines changed

6 files changed

+335
-2
lines changed

docs/references.bib

Lines changed: 9 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -178,6 +178,15 @@ @inproceedings{lewis2019
178178
keywords = {Hensel's lemma, Lean, formal proof, p-adic},
179179
}
180180

181+
@misc{pöschel2017siegelsternberg,
182+
title={On the Siegel-Sternberg linearization theorem},
183+
author={Jürgen Pöschel},
184+
year={2017},
185+
eprint={1702.03691},
186+
archivePrefix={arXiv},
187+
primaryClass={math.DS}
188+
}
189+
181190
@book {riehl2017,
182191
AUTHOR = {Riehl, Emily},
183192
TITLE = {Category theory in context},

src/algebra/geom_sum.lean

Lines changed: 4 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -272,6 +272,10 @@ theorem geom_sum_Ico [division_ring α] {x : α} (hx : x ≠ 1) {m n : ℕ} (hmn
272272
by simp only [sum_Ico_eq_sub _ hmn, (geom_series_def _ _).symm, geom_sum hx, div_sub_div_same,
273273
sub_sub_sub_cancel_right]
274274

275+
theorem geom_sum_Ico' [division_ring α] {x : α} (hx : x ≠ 1) {m n : ℕ} (hmn : m ≤ n) :
276+
∑ i in finset.Ico m n, x ^ i = (x ^ m - x ^ n) / (1 - x) :=
277+
by { simp only [geom_sum_Ico hx hmn], convert neg_div_neg_eq (x^m - x^n) (1-x); abel }
278+
275279
lemma geom_sum_inv [division_ring α] {x : α} (hx1 : x ≠ 1) (hx0 : x ≠ 0) (n : ℕ) :
276280
(geom_series x⁻¹ n) = (x - 1)⁻¹ * (x - x⁻¹ ^ n * x) :=
277281
have h₁ : x⁻¹ ≠ 1, by rwa [inv_eq_one_div, ne.def, div_eq_iff_mul_eq hx0, one_mul],

src/analysis/analytic/basic.lean

Lines changed: 12 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -182,6 +182,18 @@ lemma nnnorm_mul_pow_le_of_lt_radius (p : formal_multilinear_series 𝕜 E F) {r
182182
let ⟨C, hC, hp⟩ := p.norm_mul_pow_le_of_lt_radius h
183183
in ⟨⟨C, hC.lt.le⟩, hC, by exact_mod_cast hp⟩
184184

185+
/-- If the radius of `p` is positive, then `∥pₙ∥` grows at most geometrically. -/
186+
lemma le_mul_pow_of_radius_pos (p : formal_multilinear_series 𝕜 E F) (h : 0 < p.radius) :
187+
∃ C r (hC : 0 < C) (hr : 0 < r), ∀ n, ∥p n∥ ≤ C * r ^ n :=
188+
begin
189+
rcases ennreal.lt_iff_exists_nnreal_btwn.1 h with ⟨r, r0, rlt⟩,
190+
have rpos : 0 < (r : ℝ), by simp [ennreal.coe_pos.1 r0],
191+
rcases norm_le_div_pow_of_pos_of_lt_radius p rpos rlt with ⟨C, Cpos, hCp⟩,
192+
refine ⟨C, r ⁻¹, Cpos, by simp [rpos], λ n, _⟩,
193+
convert hCp n,
194+
exact inv_pow' _ _,
195+
end
196+
185197
/-- The radius of the sum of two formal series is at least the minimum of their two radii. -/
186198
lemma min_radius_le_radius_add (p q : formal_multilinear_series 𝕜 E F) :
187199
min p.radius q.radius ≤ (p + q).radius :=

src/analysis/analytic/composition.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -457,7 +457,7 @@ begin
457457
... ≤ Cq : hCq _,
458458
have B,
459459
calc ((∏ i, nnnorm (p (c.blocks_fun i))) * rp ^ n)
460-
∏ i, nnnorm (p (c.blocks_fun i)) * rp ^ c.blocks_fun i :
460+
= ∏ i, nnnorm (p (c.blocks_fun i)) * rp ^ c.blocks_fun i :
461461
by simp only [finset.prod_mul_distrib, finset.prod_pow_eq_pow_sum, c.sum_blocks_fun]
462462
... ≤ ∏ i : fin c.length, Cp : finset.prod_le_prod' (λ i _, hCp _)
463463
... = Cp ^ c.length : by simp

src/analysis/analytic/inverse.lean

Lines changed: 268 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -10,7 +10,7 @@ import analysis.analytic.composition
1010
# Inverse of analytic functions
1111
1212
We construct the left and right inverse of a formal multilinear series with invertible linear term,
13-
and we prove that they coincide.
13+
we prove that they coincide and study their properties (notably convergence).
1414
1515
## Main statements
1616
@@ -21,6 +21,8 @@ and we prove that they coincide.
2121
* `p.left_inv_comp` says that `p.left_inv i` is indeed a left inverse to `p` when `p₁ = i`.
2222
* `p.right_inv_comp` says that `p.right_inv i` is indeed a right inverse to `p` when `p₁ = i`.
2323
* `p.left_inv_eq_right_inv`: the two inverses coincide.
24+
* `p.radius_right_inv_pos_of_radius_pos`: if a power series has a positive radius of convergence,
25+
then so does its inverse.
2426
2527
-/
2628

@@ -268,4 +270,269 @@ left_inv p i = left_inv p.remove_zero i : by rw left_inv_remove_zero
268270
... = right_inv p.remove_zero i : by { apply left_inv_eq_right_inv_aux; simp [h] }
269271
... = right_inv p i : by rw right_inv_remove_zero
270272

273+
/-!
274+
### Convergence of the inverse of a power series
275+
276+
Assume that `p` is a convergent multilinear series, and let `q` be its (left or right) inverse.
277+
Using the left-inverse formula gives
278+
$$
279+
q_n = - (p_1)^{-n} \sum_{k=0}^{n-1} \sum_{i_1 + \dotsc + i_k = n} q_k (p_{i_1}, \dotsc, p_{i_k}).
280+
$$
281+
Assume for simplicity that we are in dimension `1` and `p₁ = 1`. In the formula for `qₙ`, the term
282+
`q_{n-1}` appears with a multiplicity of `n-1` (choosing the index `i_j` for which `i_j = 2` while
283+
all the other indices are equal to `1`), which indicates that `qₙ` might grow like `n!`. This is
284+
bad for summability properties.
285+
286+
It turns out that the right-inverse formula is better behaved, and should instead be used for this
287+
kind of estimate. It reads
288+
$$
289+
q_n = - (p_1)^{-1} \sum_{k=2}^n \sum_{i_1 + \dotsc + i_k = n} p_k (q_{i_1}, \dotsc, q_{i_k}).
290+
$$
291+
Here, `q_{n-1}` can only appear in the term with `k = 2`, and it only appears twice, so there is
292+
hope this formula can lead to an at most geometric behavior.
293+
294+
Let `Qₙ = ∥qₙ∥`. Bounding `∥pₖ∥` with `C r^k` gives an inequality
295+
$$
296+
Q_n ≤ C' \sum_{k=2}^n r^k \sum_{i_1 + \dotsc + i_k = n} Q_{i_1} \dotsm Q_{i_k}.
297+
$$
298+
299+
This formula is not enough to prove by naive induction on `n` a bound of the form `Qₙ ≤ D R^n`.
300+
However, assuming that the inequality above were an equality, one could get a formula for the
301+
generating series of the `Qₙ`:
302+
303+
$$
304+
\begin{align*}
305+
Q(z) & := \sum Q_n z^n = Q_1 z + C' \sum_{2 \leq k \leq n} \sum_{i_1 + \dotsc + i_k = n}
306+
(r z^{i_1} Q_{i_1}) \dotsm (r z^{i_k} Q_{i_k})
307+
\\& = Q_1 z + C' \sum_{k = 2}^\infty (\sum_{i_1 \geq 1} r z^{i_1} Q_{i_1})
308+
\dotsm (\sum_{i_k \geq 1} r z^{i_k} Q_{i_k})
309+
\\& = Q_1 z + C' \sum_{k = 2}^\infty (r Q(z))^k
310+
= Q_1 z + C' (r Q(z))^2 / (1 - r Q(z)).
311+
\end{align*}
312+
$$
313+
314+
One can solve this formula explicitly. The solution is analytic in a neighborhood of `0` in `ℂ`,
315+
hence its coefficients grow at most geometrically (by a contour integral argument), and therefore
316+
the original `Qₙ`, which are bounded by these ones, are also at most geometric.
317+
318+
This classical argument is not really satisfactory, as it requires an a priori bound on a complex
319+
analytic function. Another option would be to compute explicitly its terms (with binomial
320+
coefficients) to obtain an explicit geometric bound, but this would be very painful.
321+
322+
Instead, we will use the above intuition, but in a slightly different form, with finite sums and an
323+
induction. I learnt this trick in [pöschel2017siegelsternberg]. Let
324+
$S_n = \sum_{k=1}^n Q_k a^k$ (where `a` is a positive real parameter to be chosen suitably small).
325+
The above computation but with finite sums shows that
326+
327+
$$
328+
S_n \leq Q_1 a + C' \sum_{k=2}^n (r S_{n-1})^k.
329+
$$
330+
331+
In particular, $S_n \leq Q_1 a + C' (r S_{n-1})^2 / (1- r S_{n-1})$.
332+
Assume that $S_{n-1} \leq K a$, where `K > Q₁` is fixed and `a` is small enough so that
333+
`r K a ≤ 1/2` (to control the denominator). Then this equation gives a bound
334+
$S_n \leq Q_1 a + 2 C' r^2 K^2 a^2$.
335+
If `a` is small enough, this is bounded by `K a` as the second term is quadratic in `a`, and
336+
therefore negligible.
337+
338+
By induction, we deduce `Sₙ ≤ K a` for all `n`, which gives in particular the fact that `aⁿ Qₙ`
339+
remains bounded.
340+
-/
341+
342+
/-- First technical lemma to control the growth of coefficients of the inverse. Bound the explicit
343+
expression for `∑_{k<n+1} aᵏ Qₖ` in terms of a sum of powers of the same sum one step before,
344+
in a general abstract setup. -/
345+
lemma radius_right_inv_pos_of_radius_pos_aux1
346+
(n : ℕ) (p : ℕ → ℝ) (hp : ∀ k, 0 ≤ p k) {r a : ℝ} (hr : 0 ≤ r) (ha : 0 ≤ a) :
347+
∑ k in Ico 2 (n + 1), a ^ k *
348+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
349+
r ^ c.length * ∏ j, p (c.blocks_fun j))
350+
≤ ∑ j in Ico 2 (n + 1), r ^ j * (∑ k in Ico 1 n, a ^ k * p k) ^ j :=
351+
calc
352+
∑ k in Ico 2 (n + 1), a ^ k *
353+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
354+
r ^ c.length * ∏ j, p (c.blocks_fun j))
355+
= ∑ k in Ico 2 (n + 1),
356+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
357+
∏ j, r * (a ^ (c.blocks_fun j) * p (c.blocks_fun j))) :
358+
begin
359+
simp_rw [mul_sum],
360+
apply sum_congr rfl (λ k hk, _),
361+
apply sum_congr rfl (λ c hc, _),
362+
rw [prod_mul_distrib, prod_mul_distrib, prod_pow_eq_pow_sum, composition.sum_blocks_fun,
363+
prod_const, card_fin],
364+
ring,
365+
end
366+
... ≤ ∑ d in comp_partial_sum_target 2 (n + 1) n,
367+
∏ (j : fin d.2.length), r * (a ^ d.2.blocks_fun j * p (d.2.blocks_fun j)) :
368+
begin
369+
rw sum_sigma',
370+
refine sum_le_sum_of_subset_of_nonneg _ (λ x hx1 hx2,
371+
prod_nonneg (λ j hj, mul_nonneg hr (mul_nonneg (pow_nonneg ha _) (hp _)))),
372+
rintros ⟨k, c⟩ hd,
373+
simp only [set.mem_to_finset, Ico.mem, mem_sigma, set.mem_set_of_eq] at hd,
374+
simp only [mem_comp_partial_sum_target_iff],
375+
refine ⟨hd.2, c.length_le.trans_lt hd.1.2, λ j, _⟩,
376+
have : c ≠ composition.single k (zero_lt_two.trans_le hd.1.1),
377+
by simp [composition.eq_single_iff_length, ne_of_gt hd.2],
378+
rw composition.ne_single_iff at this,
379+
exact (this j).trans_le (nat.lt_succ_iff.mp hd.1.2)
380+
end
381+
... = ∑ e in comp_partial_sum_source 2 (n+1) n, ∏ (j : fin e.1), r * (a ^ e.2 j * p (e.2 j)) :
382+
begin
383+
symmetry,
384+
apply comp_change_of_variables_sum,
385+
rintros ⟨k, blocks_fun⟩ H,
386+
have K : (comp_change_of_variables 2 (n + 1) n ⟨k, blocks_fun⟩ H).snd.length = k, by simp,
387+
congr' 2; try { rw K },
388+
rw fin.heq_fun_iff K.symm,
389+
assume j,
390+
rw comp_change_of_variables_blocks_fun,
391+
end
392+
... = ∑ j in Ico 2 (n+1), r ^ j * (∑ k in Ico 1 n, a ^ k * p k) ^ j :
393+
begin
394+
rw [comp_partial_sum_source, ← sum_sigma' (Ico 2 (n + 1))
395+
(λ (k : ℕ), (fintype.pi_finset (λ (i : fin k), Ico 1 n) : finset (fin k → ℕ)))
396+
(λ n e, ∏ (j : fin n), r * (a ^ e j * p (e j)))],
397+
apply sum_congr rfl (λ j hj, _),
398+
simp only [← @multilinear_map.mk_pi_algebra_apply ℝ (fin j) _ _ ℝ],
399+
simp only [← multilinear_map.map_sum_finset (multilinear_map.mk_pi_algebra ℝ (fin j) ℝ)
400+
(λ k (m : ℕ), r * (a ^ m * p m))],
401+
simp only [multilinear_map.mk_pi_algebra_apply],
402+
dsimp,
403+
simp [prod_const, ← mul_sum, mul_pow],
404+
end
405+
406+
/-- Second technical lemma to control the growth of coefficients of the inverse. Bound the explicit
407+
expression for `∑_{k<n+1} aᵏ Qₖ` in terms of a sum of powers of the same sum one step before,
408+
in the specific setup we are interesting in, by reducing to the general bound in
409+
`radius_right_inv_pos_of_radius_pos_aux1`. -/
410+
lemma radius_right_inv_pos_of_radius_pos_aux2
411+
{n : ℕ} (hn : 2 ≤ n + 1) (p : formal_multilinear_series 𝕜 E F) (i : E ≃L[𝕜] F)
412+
{r a C : ℝ} (hr : 0 ≤ r) (ha : 0 ≤ a) (hC : 0 ≤ C) (hp : ∀ n, ∥p n∥ ≤ C * r ^ n) :
413+
(∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥) ≤
414+
∥(i.symm : F →L[𝕜] E)∥ * a + ∥(i.symm : F →L[𝕜] E)∥ * C * ∑ k in Ico 2 (n + 1),
415+
(r * ((∑ j in Ico 1 n, a ^ j * ∥p.right_inv i j∥))) ^ k :=
416+
let I := ∥(i.symm : F →L[𝕜] E)∥ in calc
417+
∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥
418+
= a * I + ∑ k in Ico 2 (n + 1), a ^ k * ∥p.right_inv i k∥ :
419+
by simp only [continuous_multilinear_curry_fin1_symm_apply_norm, pow_one, right_inv_coeff_one,
420+
Ico.succ_singleton, sum_singleton, ← sum_Ico_consecutive _ one_le_two hn]
421+
... = a * I + ∑ k in Ico 2 (n + 1), a ^ k *
422+
∥(i.symm : F →L[𝕜] E).comp_continuous_multilinear_map
423+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
424+
p.comp_along_composition (p.right_inv i) c)∥ :
425+
begin
426+
congr' 1,
427+
apply sum_congr rfl (λ j hj, _),
428+
rw [right_inv_coeff _ _ _ (Ico.mem.1 hj).1, norm_neg],
429+
end
430+
... ≤ a * ∥(i.symm : F →L[𝕜] E)∥ + ∑ k in Ico 2 (n + 1), a ^ k * (I *
431+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
432+
C * r ^ c.length * ∏ j, ∥p.right_inv i (c.blocks_fun j)∥)) :
433+
begin
434+
apply_rules [add_le_add, le_refl, sum_le_sum (λ j hj, _), mul_le_mul_of_nonneg_left,
435+
pow_nonneg, ha],
436+
apply (continuous_linear_map.norm_comp_continuous_multilinear_map_le _ _).trans,
437+
apply mul_le_mul_of_nonneg_left _ (norm_nonneg _),
438+
apply (norm_sum_le _ _).trans,
439+
apply sum_le_sum (λ c hc, _),
440+
apply (comp_along_composition_norm _ _ _).trans,
441+
apply mul_le_mul_of_nonneg_right (hp _),
442+
exact prod_nonneg (λ j hj, norm_nonneg _),
443+
end
444+
... = I * a + I * C * ∑ k in Ico 2 (n + 1), a ^ k *
445+
(∑ c in ({c | 1 < composition.length c}.to_finset : finset (composition k)),
446+
r ^ c.length * ∏ j, ∥p.right_inv i (c.blocks_fun j)∥) :
447+
begin
448+
simp_rw [mul_assoc C, ← mul_sum, ← mul_assoc, mul_comm _ (∥↑i.symm∥), mul_assoc, ← mul_sum,
449+
← mul_assoc, mul_comm _ C, mul_assoc, ← mul_sum],
450+
ring,
451+
end
452+
... ≤ I * a + I * C * ∑ k in Ico 2 (n+1), (r * ((∑ j in Ico 1 n, a ^ j * ∥p.right_inv i j∥))) ^ k :
453+
begin
454+
apply_rules [add_le_add, le_refl, mul_le_mul_of_nonneg_left, norm_nonneg, hC, mul_nonneg],
455+
simp_rw [mul_pow],
456+
apply radius_right_inv_pos_of_radius_pos_aux1 n (λ k, ∥p.right_inv i k∥)
457+
(λ k, norm_nonneg _) hr ha,
458+
end
459+
460+
/-- If a a formal multilinear series has a positive radius of convergence, then its right inverse
461+
also has a positive radius of convergence. -/
462+
theorem radius_right_inv_pos_of_radius_pos (p : formal_multilinear_series 𝕜 E F) (i : E ≃L[𝕜] F)
463+
(hp : 0 < p.radius) : 0 < (p.right_inv i).radius :=
464+
begin
465+
obtain ⟨C, r, Cpos, rpos, ple⟩ : ∃ C r (hC : 0 < C) (hr : 0 < r), ∀ (n : ℕ), ∥p n∥ ≤ C * r ^ n :=
466+
le_mul_pow_of_radius_pos p hp,
467+
let I := ∥(i.symm : F →L[𝕜] E)∥,
468+
-- choose `a` small enough to make sure that `∑_{k ≤ n} aᵏ Qₖ` will be controllable by
469+
-- induction
470+
obtain ⟨a, apos, ha1, ha2⟩ : ∃ a (apos : 0 < a),
471+
(2 * I * C * r^2 * (I + 1) ^ 2 * a ≤ 1) ∧ (r * (I + 1) * a ≤ 1/2),
472+
{ have : tendsto (λ a, 2 * I * C * r^2 * (I + 1) ^ 2 * a) (𝓝 0)
473+
(𝓝 (2 * I * C * r^2 * (I + 1) ^ 2 * 0)) := tendsto_const_nhds.mul tendsto_id,
474+
have A : ∀ᶠ a in 𝓝 0, 2 * I * C * r^2 * (I + 1) ^ 2 * a < 1,
475+
by { apply (tendsto_order.1 this).2, simp [zero_lt_one] },
476+
have : tendsto (λ a, r * (I + 1) * a) (𝓝 0)
477+
(𝓝 (r * (I + 1) * 0)) := tendsto_const_nhds.mul tendsto_id,
478+
have B : ∀ᶠ a in 𝓝 0, r * (I + 1) * a < 1/2,
479+
by { apply (tendsto_order.1 this).2, simp [zero_lt_one] },
480+
have C : ∀ᶠ a in 𝓝[set.Ioi (0 : ℝ)] (0 : ℝ), (0 : ℝ) < a,
481+
by { filter_upwards [self_mem_nhds_within], exact λ a ha, ha },
482+
rcases (C.and ((A.and B).filter_mono inf_le_left)).exists with ⟨a, ha⟩,
483+
exact ⟨a, ha.1, ha.2.1.le, ha.2.2.le⟩ },
484+
-- check by induction that the partial sums are suitably bounded, using the choice of `a` and the
485+
-- inductive control from Lemma `radius_right_inv_pos_of_radius_pos_aux2`.
486+
let S := λ n, ∑ k in Ico 1 n, a ^ k * ∥p.right_inv i k∥,
487+
have IRec : ∀ n, 1 ≤ n → S n ≤ (I + 1) * a,
488+
{ apply nat.le_induction,
489+
{ simp only [S],
490+
rw [Ico.eq_empty_of_le (le_refl 1), sum_empty],
491+
exact mul_nonneg (add_nonneg (norm_nonneg _) zero_le_one) apos.le },
492+
{ assume n one_le_n hn,
493+
have In : 2 ≤ n + 1, by linarith,
494+
have Snonneg : 0 ≤ S n :=
495+
sum_nonneg (λ x hx, mul_nonneg (pow_nonneg apos.le _) (norm_nonneg _)),
496+
have rSn : r * S n ≤ 1/2 := calc
497+
r * S n ≤ r * ((I+1) * a) : mul_le_mul_of_nonneg_left hn rpos.le
498+
... ≤ 1/2 : by rwa [← mul_assoc],
499+
calc S (n + 1) ≤ I * a + I * C * ∑ k in Ico 2 (n + 1), (r * S n)^k :
500+
radius_right_inv_pos_of_radius_pos_aux2 In p i rpos.le apos.le Cpos.le ple
501+
... = I * a + I * C * (((r * S n) ^ 2 - (r * S n) ^ (n + 1)) / (1 - r * S n)) :
502+
by { rw geom_sum_Ico' _ In, exact ne_of_lt (rSn.trans_lt (by norm_num)) }
503+
... ≤ I * a + I * C * ((r * S n) ^ 2 / (1/2)) :
504+
begin
505+
apply_rules [add_le_add, le_refl, mul_le_mul_of_nonneg_left, mul_nonneg, norm_nonneg,
506+
Cpos.le],
507+
refine div_le_div (pow_two_nonneg _) _ (by norm_num) (by linarith),
508+
simp only [sub_le_self_iff],
509+
apply pow_nonneg (mul_nonneg rpos.le Snonneg),
510+
end
511+
... = I * a + 2 * I * C * (r * S n) ^ 2 : by ring
512+
... ≤ I * a + 2 * I * C * (r * ((I + 1) * a)) ^ 2 :
513+
by apply_rules [add_le_add, le_refl, mul_le_mul_of_nonneg_left, mul_nonneg, norm_nonneg,
514+
Cpos.le, zero_le_two, pow_le_pow_of_le_left, rpos.le]
515+
... = (I + 2 * I * C * r^2 * (I + 1) ^ 2 * a) * a : by ring
516+
... ≤ (I + 1) * a :
517+
by apply_rules [mul_le_mul_of_nonneg_right, apos.le, add_le_add, le_refl] } },
518+
-- conclude that all coefficients satisfy `aⁿ Qₙ ≤ (I + 1) a`.
519+
let a' : nnreal := ⟨a, apos.le⟩,
520+
suffices H : (a' : ennreal) ≤ (p.right_inv i).radius,
521+
by { apply lt_of_lt_of_le _ H, exact_mod_cast apos },
522+
apply le_radius_of_bound _ ((I + 1) * a) (λ n, _),
523+
by_cases hn : n = 0,
524+
{ have : ∥p.right_inv i n∥ = ∥p.right_inv i 0∥, by congr; try { rw hn },
525+
simp only [this, norm_zero, zero_mul, right_inv_coeff_zero],
526+
apply_rules [mul_nonneg, add_nonneg, norm_nonneg, zero_le_one, apos.le] },
527+
{ have one_le_n : 1 ≤ n := bot_lt_iff_ne_bot.2 hn,
528+
calc ∥p.right_inv i n∥ * ↑a' ^ n = a ^ n * ∥p.right_inv i n∥ : mul_comm _ _
529+
... ≤ ∑ k in Ico 1 (n + 1), a ^ k * ∥p.right_inv i k∥ :
530+
begin
531+
have : ∀ k ∈ Ico 1 (n + 1), 0 ≤ a ^ k * ∥p.right_inv i k∥ :=
532+
λ k hk, mul_nonneg (pow_nonneg apos.le _) (norm_nonneg _),
533+
exact single_le_sum this (by simp [one_le_n]),
534+
end
535+
... ≤ (I + 1) * a : IRec (n + 1) (by dec_trivial) }
536+
end
537+
271538
end formal_multilinear_series

0 commit comments

Comments
 (0)