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

Commit 3a136f8

Browse files
committed
refactor(analysis/analytic/composition): extend definition, extract a lemma from a proof (#5850)
Extract a standalone lemma of the proof that the composition of two analytic functions is well-behaved, and extend a little bit the definition of the sets which are involved in the corresponding change of variables.
1 parent 74760f2 commit 3a136f8

File tree

1 file changed

+84
-66
lines changed

1 file changed

+84
-66
lines changed

src/analysis/analytic/composition.lean

Lines changed: 84 additions & 66 deletions
Original file line numberDiff line numberDiff line change
@@ -432,21 +432,23 @@ the source of the change of variables (`comp_partial_source`), its target
432432
(`comp_partial_target`) and the change of variables itself (`comp_change_of_variables`) before
433433
giving the main statement in `comp_partial_sum`. -/
434434

435+
435436
/-- Source set in the change of variables to compute the composition of partial sums of formal
436437
power series.
437438
See also `comp_partial_sum`. -/
438-
def comp_partial_sum_source (N : ℕ) : finset (Σ n, (fin n) → ℕ) :=
439-
finset.sigma (finset.range N) (λ (n : ℕ), fintype.pi_finset (λ (i : fin n), finset.Ico 1 N) : _)
439+
def comp_partial_sum_source (m M N : ℕ) : finset (Σ n, (fin n) → ℕ) :=
440+
finset.sigma (finset.Ico m M) (λ (n : ℕ), fintype.pi_finset (λ (i : fin n), finset.Ico 1 N) : _)
440441

441-
@[simp] lemma mem_comp_partial_sum_source_iff (N : ℕ) (i : Σ n, (fin n) → ℕ) :
442-
i ∈ comp_partial_sum_source N ↔ i.1 < N ∧ ∀ (a : fin i.1), 1 ≤ i.2 a ∧ i.2 a < N :=
443-
by simp only [comp_partial_sum_source, finset.Ico.mem,
444-
fintype.mem_pi_finset, finset.mem_sigma, finset.mem_range]
442+
@[simp] lemma mem_comp_partial_sum_source_iff (m M N : ℕ) (i : Σ n, (fin n) → ℕ) :
443+
i ∈ comp_partial_sum_source m M N ↔
444+
(m ≤ i.1 ∧ i.1 < M) ∧ ∀ (a : fin i.1), 1 ≤ i.2 a ∧ i.2 a < N :=
445+
by simp only [comp_partial_sum_source, finset.Ico.mem, fintype.mem_pi_finset, finset.mem_sigma,
446+
iff_self]
445447

446448
/-- Change of variables appearing to compute the composition of partial sums of formal
447449
power series -/
448-
def comp_change_of_variables (N : ℕ) (i : Σ n, (fin n) → ℕ) (hi : i ∈ comp_partial_sum_source N) :
449-
(Σ n, composition n) :=
450+
def comp_change_of_variables (m M N : ℕ) (i : Σ n, (fin n) → ℕ)
451+
(hi : i ∈ comp_partial_sum_source m M N) : (Σ n, composition n) :=
450452
begin
451453
rcases i with ⟨n, f⟩,
452454
rw mem_comp_partial_sum_source_iff at hi,
@@ -456,18 +458,18 @@ begin
456458
end
457459

458460
@[simp] lemma comp_change_of_variables_length
459-
(N : ℕ) {i : Σ n, (fin n) → ℕ} (hi : i ∈ comp_partial_sum_source N) :
460-
composition.length (comp_change_of_variables N i hi).2 = i.1 :=
461+
(m M N : ℕ) {i : Σ n, (fin n) → ℕ} (hi : i ∈ comp_partial_sum_source m M N) :
462+
composition.length (comp_change_of_variables m M N i hi).2 = i.1 :=
461463
begin
462464
rcases i with ⟨k, blocks_fun⟩,
463465
dsimp [comp_change_of_variables],
464466
simp only [composition.length, map_of_fn, length_of_fn]
465467
end
466468

467469
lemma comp_change_of_variables_blocks_fun
468-
(N : ℕ) {i : Σ n, (fin n) → ℕ} (hi : i ∈ comp_partial_sum_source N) (j : fin i.1) :
469-
(comp_change_of_variables N i hi).2.blocks_fun
470-
⟨j, (comp_change_of_variables_length N hi).symm ▸ j.2⟩ = i.2 j :=
470+
(m M N : ℕ) {i : Σ n, (fin n) → ℕ} (hi : i ∈ comp_partial_sum_source m M N) (j : fin i.1) :
471+
(comp_change_of_variables m M N i hi).2.blocks_fun
472+
⟨j, (comp_change_of_variables_length m M N hi).symm ▸ j.2⟩ = i.2 j :=
471473
begin
472474
rcases i with ⟨n, f⟩,
473475
dsimp [composition.blocks_fun, composition.blocks, comp_change_of_variables],
@@ -478,12 +480,12 @@ end
478480

479481
/-- Target set in the change of variables to compute the composition of partial sums of formal
480482
power series, here given a a set. -/
481-
def comp_partial_sum_target_set (N : ℕ) : set (Σ n, composition n) :=
482-
{i | (i.2.length < N) ∧ (∀ (j : fin i.2.length), i.2.blocks_fun j < N)}
483+
def comp_partial_sum_target_set (m M N : ℕ) : set (Σ n, composition n) :=
484+
{i | (m ≤ i.2.length) ∧ (i.2.length < M) ∧ (∀ (j : fin i.2.length), i.2.blocks_fun j < N)}
483485

484486
lemma comp_partial_sum_target_subset_image_comp_partial_sum_source
485-
(N : ℕ) (i : Σ n, composition n) (hi : i ∈ comp_partial_sum_target_set N) :
486-
∃ j (hj : j ∈ comp_partial_sum_source N), i = comp_change_of_variables N j hj :=
487+
(m M N : ℕ) (i : Σ n, composition n) (hi : i ∈ comp_partial_sum_target_set m M N) :
488+
∃ j (hj : j ∈ comp_partial_sum_source m M N), i = comp_change_of_variables m M N j hj :=
487489
begin
488490
rcases i with ⟨n, c⟩,
489491
refine ⟨⟨c.length, c.blocks_fun⟩, _, _⟩,
@@ -493,25 +495,69 @@ begin
493495
{ dsimp [comp_change_of_variables],
494496
rw composition.sigma_eq_iff_blocks_eq,
495497
simp only [composition.blocks_fun, composition.blocks, subtype.coe_eta, nth_le_map'],
496-
conv_lhs { rw ← of_fn_nth_le c.blocks },
497-
simp only [fin.val_eq_coe], refl, /- where does this fin.val come from? -/ }
498+
conv_lhs { rw ← of_fn_nth_le c.blocks }, refl, /- where does this fin.val come from? -/ }
498499
end
499500

500501
/-- Target set in the change of variables to compute the composition of partial sums of formal
501502
power series, here given a a finset.
502503
See also `comp_partial_sum`. -/
503-
def comp_partial_sum_target (N : ℕ) : finset (Σ n, composition n) :=
504+
def comp_partial_sum_target (m M N : ℕ) : finset (Σ n, composition n) :=
504505
set.finite.to_finset $ (finset.finite_to_set _).dependent_image
505-
(comp_partial_sum_target_subset_image_comp_partial_sum_source N)
506+
(comp_partial_sum_target_subset_image_comp_partial_sum_source m M N)
506507

507-
@[simp] lemma mem_comp_partial_sum_target_iff {N : ℕ} {a : Σ n, composition n} :
508-
a ∈ comp_partial_sum_target N ↔ a.2.length < N ∧ (∀ (j : fin a.2.length), a.2.blocks_fun j < N) :=
508+
@[simp] lemma mem_comp_partial_sum_target_iff {m M N : ℕ} {a : Σ n, composition n} :
509+
a ∈ comp_partial_sum_target m M N ↔
510+
m ≤ a.2.length ∧ a.2.length < M ∧ (∀ (j : fin a.2.length), a.2.blocks_fun j < N) :=
509511
by simp [comp_partial_sum_target, comp_partial_sum_target_set]
510512

513+
/-- `comp_change_of_variables m M N` is a bijection between `comp_partial_sum_source m M N`
514+
and `comp_partial_sum_target m M N`, yielding equal sums for functions that correspond to each
515+
other under the bijection. As `comp_change_of_variables m M N` is a dependent function, stating
516+
that it is a bijection is not directly possible, but the consequence on sums can be stated
517+
more easily. -/
518+
lemma comp_change_of_variables_sum {α : Type*} [add_comm_monoid α] (m M N : ℕ)
519+
(f : (Σ (n : ℕ), fin n → ℕ) → α) (g : (Σ n, composition n) → α)
520+
(h : ∀ e (he : e ∈ comp_partial_sum_source m M N), f e = g (comp_change_of_variables m M N e he)) :
521+
∑ e in comp_partial_sum_source m M N, f e = ∑ e in comp_partial_sum_target m M N, g e :=
522+
begin
523+
apply finset.sum_bij (comp_change_of_variables m M N),
524+
-- We should show that the correspondance we have set up is indeed a bijection
525+
-- between the index sets of the two sums.
526+
-- 1 - show that the image belongs to `comp_partial_sum_target N N`
527+
{ rintros ⟨k, blocks_fun⟩ H,
528+
rw mem_comp_partial_sum_source_iff at H,
529+
simp only [mem_comp_partial_sum_target_iff, composition.length, composition.blocks, H.left,
530+
map_of_fn, length_of_fn, true_and, comp_change_of_variables],
531+
assume j,
532+
simp only [composition.blocks_fun, (H.right _).right, nth_le_of_fn'] },
533+
-- 2 - show that the composition gives the `comp_along_composition` application
534+
{ rintros ⟨k, blocks_fun⟩ H,
535+
rw h },
536+
-- 3 - show that the map is injective
537+
{ rintros ⟨k, blocks_fun⟩ ⟨k', blocks_fun'⟩ H H' heq,
538+
obtain rfl : k = k',
539+
{ have := (comp_change_of_variables_length m M N H).symm,
540+
rwa [heq, comp_change_of_variables_length] at this, },
541+
congr,
542+
funext i,
543+
calc blocks_fun i = (comp_change_of_variables m M N _ H).2.blocks_fun _ :
544+
(comp_change_of_variables_blocks_fun m M N H i).symm
545+
... = (comp_change_of_variables m M N _ H').2.blocks_fun _ :
546+
begin
547+
apply composition.blocks_fun_congr; try { rw heq },
548+
refl
549+
end
550+
... = blocks_fun' i : comp_change_of_variables_blocks_fun m M N H' i },
551+
-- 4 - show that the map is surjective
552+
{ assume i hi,
553+
apply comp_partial_sum_target_subset_image_comp_partial_sum_source m M N i,
554+
simpa [comp_partial_sum_target] using hi }
555+
end
556+
511557
/-- The auxiliary set corresponding to the composition of partial sums asymptotically contains
512558
all possible compositions. -/
513559
lemma comp_partial_sum_target_tendsto_at_top :
514-
tendsto comp_partial_sum_target at_top at_top :=
560+
tendsto (λ N, comp_partial_sum_target 0 N N) at_top at_top :=
515561
begin
516562
apply monotone.tendsto_at_top_finset,
517563
{ assume m n hmn a ha,
@@ -521,7 +567,7 @@ begin
521567
simp only [mem_comp_partial_sum_target_iff],
522568
obtain ⟨n, hn⟩ : bdd_above ↑(finset.univ.image (λ (i : fin c.length), c.blocks_fun i)) :=
523569
finset.bdd_above _,
524-
refine ⟨max n c.length + 1, lt_of_le_of_lt (le_max_right n c.length) (lt_add_one _),
570+
refine ⟨max n c.length + 1, bot_le, lt_of_le_of_lt (le_max_right n c.length) (lt_add_one _),
525571
λ j, lt_of_le_of_lt (le_trans _ (le_max_left _ _)) (lt_add_one _)⟩,
526572
apply hn,
527573
simp only [finset.mem_image_of_mem, finset.mem_coe, finset.mem_univ] }
@@ -533,52 +579,24 @@ compositions in `comp_partial_sum_target N`. This is precisely the motivation fo
533579
lemma comp_partial_sum
534580
(q : formal_multilinear_series 𝕜 F G) (p : formal_multilinear_series 𝕜 E F) (N : ℕ) (z : E) :
535581
q.partial_sum N (∑ i in finset.Ico 1 N, p i (λ j, z)) =
536-
∑ i in comp_partial_sum_target N, q.comp_along_composition_multilinear p i.2 (λ j, z) :=
582+
∑ i in comp_partial_sum_target 0 N N, q.comp_along_composition p i.2 (λ j, z) :=
537583
begin
538584
-- we expand the composition, using the multilinearity of `q` to expand along each coordinate.
539585
suffices H : ∑ n in finset.range N, ∑ r in fintype.pi_finset (λ (i : fin n), finset.Ico 1 N),
540586
q n (λ (i : fin n), p (r i) (λ j, z)) =
541-
∑ i in comp_partial_sum_target N, q.comp_along_composition_multilinear p i.2 (λ j, z),
587+
∑ i in comp_partial_sum_target 0 N N, q.comp_along_composition p i.2 (λ j, z),
542588
by simpa only [formal_multilinear_series.partial_sum,
543589
continuous_multilinear_map.map_sum_finset] using H,
544-
-- rewrite the first sum as a big sum over a sigma type
545-
rw [finset.sum_sigma'],
546-
-- show that the two sums correspond to each other by reindexing the variables.
547-
apply finset.sum_bij (comp_change_of_variables N),
548-
-- To conclude, we should show that the correspondance we have set up is indeed a bijection
549-
-- between the index sets of the two sums.
550-
-- 1 - show that the image belongs to `comp_partial_sum_target N`
551-
{ rintros ⟨k, blocks_fun⟩ H,
552-
rw mem_comp_partial_sum_source_iff at H,
553-
simp only [mem_comp_partial_sum_target_iff, composition.length, composition.blocks, H.left,
554-
map_of_fn, length_of_fn, true_and, comp_change_of_variables],
555-
assume j,
556-
simp only [composition.blocks_fun, (H.right _).right, nth_le_of_fn'] },
557-
-- 2 - show that the composition gives the `comp_along_composition` application
558-
{ rintros ⟨k, blocks_fun⟩ H,
559-
apply congr _ (comp_change_of_variables_length N H).symm,
560-
intros,
561-
rw ← comp_change_of_variables_blocks_fun N H,
562-
refl },
563-
-- 3 - show that the map is injective
564-
{ rintros ⟨k, blocks_fun⟩ ⟨k', blocks_fun'⟩ H H' heq,
565-
obtain rfl : k = k',
566-
{ have := (comp_change_of_variables_length N H).symm,
567-
rwa [heq, comp_change_of_variables_length] at this, },
568-
congr,
569-
funext i,
570-
calc blocks_fun i = (comp_change_of_variables N _ H).2.blocks_fun _ :
571-
(comp_change_of_variables_blocks_fun N H i).symm
572-
... = (comp_change_of_variables N _ H').2.blocks_fun _ :
573-
begin
574-
apply composition.blocks_fun_congr; try { rw heq },
575-
refl
576-
end
577-
... = blocks_fun' i : comp_change_of_variables_blocks_fun N H' i },
578-
-- 4 - show that the map is surjective
579-
{ assume i hi,
580-
apply comp_partial_sum_target_subset_image_comp_partial_sum_source N i,
581-
simpa [comp_partial_sum_target] using hi }
590+
-- rewrite the first sum as a big sum over a sigma type, in the finset
591+
-- `comp_partial_sum_target 0 N N`
592+
rw [finset.range_eq_Ico, finset.sum_sigma'],
593+
-- use `comp_change_of_variables_sum`, saying that this change of variables respects sums
594+
apply comp_change_of_variables_sum 0 N N,
595+
rintros ⟨k, blocks_fun⟩ H,
596+
apply congr _ (comp_change_of_variables_length 0 N N H).symm,
597+
intros,
598+
rw ← comp_change_of_variables_blocks_fun 0 N N H,
599+
refl
582600
end
583601

584602
end formal_multilinear_series
@@ -662,7 +680,7 @@ begin
662680
-- `g (f (x + y))`. As this sum is exactly the composition of the partial sum, this is a direct
663681
-- consequence of the second step
664682
have C : tendsto (λ n,
665-
∑ i in comp_partial_sum_target n, q.comp_along_composition_multilinear p i.2 (λ j, y))
683+
∑ i in comp_partial_sum_target 0 n n, q.comp_along_composition_multilinear p i.2 (λ j, y))
666684
at_top (𝓝 (g (f (x + y)))),
667685
by simpa [comp_partial_sum] using B,
668686
-- Fourth step: the sum over all compositions is `g (f (x + y))`. This follows from the

0 commit comments

Comments
 (0)