Skip to content

Commit 1e908eb

Browse files
committed
feat(FaaDiBruno): add estimates on the difference between taylor series (#31937)
Prove an estimate on `(f ∘ g)^{(n)}` that can be used to show, e.g., that the composition of two `C^{n+(α)}` functions is a `C^{n+(α)}` function.
1 parent 98ea690 commit 1e908eb

File tree

2 files changed

+101
-0
lines changed

2 files changed

+101
-0
lines changed

Mathlib/Analysis/Calculus/ContDiff/FaaDiBruno.lean

Lines changed: 93 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -835,6 +835,26 @@ theorem norm_compAlongOrderedFinpartitionL_le :
835835
‖c.compAlongOrderedFinpartitionL 𝕜 E F G‖ ≤ 1 :=
836836
MultilinearMap.mkContinuousLinear_norm_le _ zero_le_one _
837837

838+
theorem norm_compAlongOrderedFinpartitionL_apply_le (f : F [×c.length]→L[𝕜] G) :
839+
‖c.compAlongOrderedFinpartitionL 𝕜 E F G f‖ ≤ ‖f‖ :=
840+
(ContinuousLinearMap.le_of_opNorm_le _ c.norm_compAlongOrderedFinpartitionL_le f).trans_eq
841+
(one_mul _)
842+
843+
theorem norm_compAlongOrderedFinpartition_sub_compAlongOrderedFinpartition_le
844+
(f₁ f₂ : F [×c.length]→L[𝕜] G) (g₁ g₂ : ∀ i, E [×c.partSize i]→L[𝕜] F) :
845+
‖c.compAlongOrderedFinpartition f₁ g₁ - c.compAlongOrderedFinpartition f₂ g₂‖ ≤
846+
‖f₁‖ * c.length * max ‖g₁‖ ‖g₂‖ ^ (c.length - 1) * ‖g₁ - g₂‖ + ‖f₁ - f₂‖ * ∏ i, ‖g₂ i‖ := calc
847+
_ ≤ ‖c.compAlongOrderedFinpartition f₁ g₁ - c.compAlongOrderedFinpartition f₁ g₂‖ +
848+
‖c.compAlongOrderedFinpartition f₁ g₂ - c.compAlongOrderedFinpartition f₂ g₂‖ :=
849+
norm_sub_le_norm_sub_add_norm_sub ..
850+
_ ≤ ‖f₁‖ * c.length * max ‖g₁‖ ‖g₂‖ ^ (c.length - 1) * ‖g₁ - g₂‖ + ‖f₁ - f₂‖ * ∏ i, ‖g₂ i‖ := by
851+
gcongr ?_ + ?_
852+
· refine ((c.compAlongOrderedFinpartitionL 𝕜 E F G f₁).norm_image_sub_le g₁ g₂).trans ?_
853+
simp only [Fintype.card_fin]
854+
gcongr
855+
apply norm_compAlongOrderedFinpartitionL_apply_le
856+
· exact c.norm_compAlongOrderedFinpartition_le (f₁ - f₂) g₂
857+
838858
end OrderedFinpartition
839859

840860
/-! ### The Faa di Bruno formula -/
@@ -876,6 +896,79 @@ protected noncomputable def taylorComp
876896
FormalMultilinearSeries 𝕜 E G :=
877897
fun n ↦ ∑ c : OrderedFinpartition n, q.compAlongOrderedFinpartition p c
878898

899+
/-- An upper estimate (in terms of `Asymptotics.IsBigO`)
900+
on the difference between two compositions of Taylor series.
901+
902+
Let `p₁`, `p₂`, `q₁`, `q₂` be four families of formal multilinear series
903+
depending on a parameter `a`.
904+
Suppose that the norms of `(p₁ · k)`, `(q₁ · k)`, and `(q₂ · k)` are bounded along a filter `l`
905+
for all `k ≤ n`.
906+
Also, suppose that $p₁(a, k) - p₂(a, k) = O(f(a))$, $q₁(a, k) - q₂(a, k) = O(f(a))$
907+
along `l` for all `k ≤ n`.
908+
Then the difference between `n`th terms of `(p₁ a).taylorComp (q₁ a)` and `(p₂ a).taylorComp (q₂ a)`
909+
is `O(f(a))` too.
910+
911+
This lemma can be used, e.g., to show that the composition of two $C^{k+α}$ functions
912+
is a $C^{k+α}$ function. -/
913+
theorem taylorComp_sub_taylorComp_isBigO
914+
{α H : Type*} [NormedAddCommGroup H] {l : Filter α} {p₁ p₂ : α → FormalMultilinearSeries 𝕜 F G}
915+
{q₁ q₂ : α → FormalMultilinearSeries 𝕜 E F} {f : α → H} {n : ℕ}
916+
(hp_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖p₁ · k‖))
917+
(hpf : ∀ k ≤ n, (fun a ↦ p₁ a k - p₂ a k) =O[l] f)
918+
(hq₁_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖q₁ · k‖))
919+
(hq₂_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖q₂ · k‖))
920+
(hqf : ∀ k ≤ n, (fun a ↦ q₁ a k - q₂ a k) =O[l] f) :
921+
(fun a ↦ (p₁ a).taylorComp (q₁ a) n - (p₂ a).taylorComp (q₂ a) n) =O[l] f := by
922+
simp only [FormalMultilinearSeries.taylorComp, ← Finset.sum_sub_distrib]
923+
refine .sum fun c _ ↦ ?_
924+
refine .trans (.of_norm_le fun _ ↦
925+
c.norm_compAlongOrderedFinpartition_sub_compAlongOrderedFinpartition_le ..) ?_
926+
refine .add ?_ ?_
927+
· have H₁ : (p₁ · c.length) =O[l] (1 : α → ℝ) := (hp_bdd _ c.length_le).isBigO_one ℝ
928+
have H₂ : ∀ m, (q₁ · (c.partSize m)) =O[l] (1 : α → ℝ) := fun m ↦
929+
(hq₁_bdd _ <| c.partSize_le _).isBigO_one ℝ
930+
have H₃ : ∀ m, (q₂ · (c.partSize m)) =O[l] (1 : α → ℝ) := fun m ↦
931+
(hq₂_bdd _ <| c.partSize_le _).isBigO_one ℝ
932+
have H₄ : ∀ m, (fun a ↦ q₁ a (c.partSize m) - q₂ a (c.partSize m)) =O[l] f := fun m ↦
933+
hqf _ <| c.partSize_le _
934+
rw [← Asymptotics.isBigO_pi] at H₂ H₃ H₄
935+
have H₅ := ((H₂.prod_left H₃).norm_left.pow (c.length - 1)).mul H₄.norm_norm
936+
simpa [mul_assoc] using H₁.norm_left.mul <| H₅.const_mul_left c.length
937+
· have H₁ : (fun a ↦ p₁ a c.length - p₂ a c.length) =O[l] f := hpf _ c.length_le
938+
have H₂ : ∀ i, (q₂ · (c.partSize i)) =O[l] (1 : α → ℝ) := fun i ↦
939+
(hq₂_bdd _ <| c.partSize_le i).isBigO_one ℝ
940+
simpa using H₁.norm_norm.mul <| .finsetProd fun i _ ↦ (H₂ i).norm_left
941+
942+
/-- An upper estimate (in terms of `Asymptotics.IsLittleO`)
943+
on the difference between two compositions of Taylor series.
944+
945+
Let `p₁`, `p₂`, `q₁`, `q₂` be four families of formal multilinear series
946+
depending on a parameter `a`.
947+
Suppose that the norms of `(p₁ · k)`, `(q₁ · k)`, and `(q₂ · k)` are bounded along a filter `l`
948+
for all `k ≤ n`.
949+
Also, suppose that $p₁(a, k) - p₂(a, k) = o(f(a))$, $q₁(a, k) - q₂(a, k) = o(f(a))$
950+
along `l` for all `k ≤ n`.
951+
Then the difference between `n`th terms of `(p₁ a).taylorComp (q₁ a)` and `(p₂ a).taylorComp (q₂ a)`
952+
is `o(f(a))` too.
953+
-/
954+
theorem taylorComp_sub_taylorComp_isLittleO
955+
{α H : Type*} [NormedAddCommGroup H] {l : Filter α} {p₁ p₂ : α → FormalMultilinearSeries 𝕜 F G}
956+
{q₁ q₂ : α → FormalMultilinearSeries 𝕜 E F} {f : α → H} {n : ℕ}
957+
(hp_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖p₁ · k‖))
958+
(hpf : ∀ k ≤ n, (fun a ↦ p₁ a k - p₂ a k) =o[l] f)
959+
(hq₁_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖q₁ · k‖))
960+
(hq₂_bdd : ∀ k ≤ n, l.IsBoundedUnder (· ≤ ·) (‖q₂ · k‖))
961+
(hqf : ∀ k ≤ n, (fun a ↦ q₁ a k - q₂ a k) =o[l] f) :
962+
(fun a ↦ (p₁ a).taylorComp (q₁ a) n - (p₂ a).taylorComp (q₂ a) n) =o[l] f := calc
963+
_ =O[l] fun a ↦ (fun k : Fin (n + 1) ↦ p₁ a k - p₂ a k,
964+
fun k : Fin (n + 1) ↦ q₁ a k - q₂ a k) := by
965+
refine taylorComp_sub_taylorComp_isBigO hp_bdd ?_ hq₁_bdd hq₂_bdd ?_
966+
all_goals simp only [← Nat.lt_succ_iff, Nat.forall_lt_iff_fin, ← Asymptotics.isBigO_pi]
967+
exacts [Asymptotics.isBigO_fst_prod, Asymptotics.isBigO_snd_prod]
968+
_ =o[l] f :=
969+
.prod_left (Asymptotics.isLittleO_pi.2 fun k ↦ hpf k (by grind))
970+
(Asymptotics.isLittleO_pi.2 fun k ↦ hqf k (by grind))
971+
879972
end FormalMultilinearSeries
880973

881974
theorem analyticOn_taylorComp

Mathlib/Data/Fin/Basic.lean

Lines changed: 8 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -33,6 +33,14 @@ open Fin Nat Function
3333

3434
attribute [simp] Fin.succ_ne_zero Fin.castSucc_lt_last
3535

36+
theorem Nat.forall_lt_iff_fin {n : ℕ} {p : ∀ k, k < n → Prop} :
37+
(∀ k hk, p k hk) ↔ ∀ k : Fin n, p k k.is_lt :=
38+
.symm <| Fin.forall_iff
39+
40+
theorem Nat.exists_lt_iff_fin {n : ℕ} {p : ∀ k, k < n → Prop} :
41+
(∃ k hk, p k hk) ↔ ∃ k : Fin n, p k k.is_lt :=
42+
.symm <| Fin.exists_iff
43+
3644
/-- Elimination principle for the empty set `Fin 0`, dependent version. -/
3745
def finZeroElim {α : Fin 0Sort*} (x : Fin 0) : α x :=
3846
x.elim0

0 commit comments

Comments
 (0)