Documentation

Mathlib.Analysis.Calculus.ContDiff.RestrictScalars

Restricting Scalars in Iterated FrΓ©chet Derivatives #

This file establishes standard theorems on restriction of scalars for iterated FrΓ©chet derivatives, comparing iterated derivatives with respect to a field π•œ' to iterated derivatives with respect to a subfield π•œ βŠ† π•œ'. The results are analogous to those found in Mathlib.Analysis.Calculus.FDeriv.RestrictScalars.

theorem fderivWithin_restrictScalars_comp {π•œ : Type u_1} {π•œ' : Type u_2} [NontriviallyNormedField π•œ] [NontriviallyNormedField π•œ'] [NormedAlgebra π•œ π•œ'] {E : Type u_3} [NormedAddCommGroup E] [NormedSpace π•œ E] [NormedSpace π•œ' E] [IsScalarTower π•œ π•œ' E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] [NormedSpace π•œ' F] [IsScalarTower π•œ π•œ' F] {x : E} {n : β„•} {s : Set E} {Ο† : E β†’ ContinuousMultilinearMap π•œ' (fun (x : Fin n) => E) F} (h : DifferentiableWithinAt π•œ' Ο† s x) (hs : UniqueDiffWithinAt π•œ s x) :

Derviation rule for compositions of scalar restriction with continuous multilinear maps.

theorem ContDiffWithinAt.restrictScalars_iteratedFDerivWithin_eventuallyEq {π•œ : Type u_1} {π•œ' : Type u_2} [NontriviallyNormedField π•œ] [NontriviallyNormedField π•œ'] [NormedAlgebra π•œ π•œ'] {E : Type u_3} [NormedAddCommGroup E] [NormedSpace π•œ E] [NormedSpace π•œ' E] [IsScalarTower π•œ π•œ' E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] [NormedSpace π•œ' F] [IsScalarTower π•œ π•œ' F] {x : E} {f : E β†’ F} {n : β„•} {s : Set E} (h : ContDiffWithinAt π•œ' (↑n) f s x) (hs : UniqueDiffOn π•œ s) (hx : x ∈ s) :

If f is n times continuously differentiable at x within s, then the nth iterated FrΓ©chet derivative within s with respect to π•œ equals scalar restriction of the nth iterated FrΓ©chet derivative within s with respect to π•œ'.

theorem ContDiffAt.restrictScalars_iteratedFDeriv_eventuallyEq {π•œ : Type u_1} {π•œ' : Type u_2} [NontriviallyNormedField π•œ] [NontriviallyNormedField π•œ'] [NormedAlgebra π•œ π•œ'] {E : Type u_3} [NormedAddCommGroup E] [NormedSpace π•œ E] [NormedSpace π•œ' E] [IsScalarTower π•œ π•œ' E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] [NormedSpace π•œ' F] [IsScalarTower π•œ π•œ' F] {x : E} {f : E β†’ F} {n : β„•} (h : ContDiffAt π•œ' (↑n) f x) :

If f is n times continuously differentiable at x, then the nth iterated FrΓ©chet derivative with respect to π•œ equals scalar restriction of the nth iterated FrΓ©chet derivative with respect to π•œ'.

theorem ContDiffAt.restrictScalars_iteratedFDeriv {π•œ : Type u_1} {π•œ' : Type u_2} [NontriviallyNormedField π•œ] [NontriviallyNormedField π•œ'] [NormedAlgebra π•œ π•œ'] {E : Type u_3} [NormedAddCommGroup E] [NormedSpace π•œ E] [NormedSpace π•œ' E] [IsScalarTower π•œ π•œ' E] {F : Type u_4} [NormedAddCommGroup F] [NormedSpace π•œ F] [NormedSpace π•œ' F] [IsScalarTower π•œ π•œ' F] {x : E} {f : E β†’ F} {n : β„•} (h : ContDiffAt π•œ' (↑n) f x) :

If f is n times continuously differentiable at x, then the nth iterated FrΓ©chet derivative with respect to π•œ equals scalar restriction of the nth iterated FrΓ©chet derivative with respect to π•œ'.