Formal multilinear series #
In this file we define FormalMultilinearSeries 𝕜 E F to be a family of n-multilinear maps for
all n, designed to model the sequence of derivatives of a function. In other files we use this
notion to define C^n functions (called contDiff in mathlib) and analytic functions.
Notations #
We use the notation E [×n]→L[𝕜] F for the space of continuous multilinear maps on E^n with
values in F. This is the space in which the n-th derivative of a function from E to F lives.
Tags #
multilinear, formal series
A formal multilinear series over a field 𝕜, from E to F, is given by a family of
multilinear maps from E^n to F for all n.
Equations
- FormalMultilinearSeries 𝕜 E F = ((n : ℕ) → ContinuousMultilinearMap 𝕜 (fun (i : Fin n) => E) F)
Instances For
Equations
- instAddCommGroupFormalMultilinearSeries = inferInstanceAs (AddCommGroup ((n : ℕ) → ContinuousMultilinearMap 𝕜 (fun (i : Fin n) => E) F))
Equations
- instInhabitedFormalMultilinearSeries = { default := 0 }
Equations
- instModuleFormalMultilinearSeriesOfContinuousConstSMulOfSMulCommClass 𝕜' = inferInstanceAs (Module 𝕜' ((n : ℕ) → ContinuousMultilinearMap 𝕜 (fun (i : Fin n) => E) F))
Cartesian product of two formal multilinear series (with the same field 𝕜 and the same source
space, but possibly different target spaces).
Equations
- p.prod q x = let n := x; (p n).prod (q n)
Instances For
Killing the zeroth coefficient in a formal multilinear series
Instances For
Convenience congruence lemma stating in a dependent setting that, if the arguments to a formal multilinear series are equal, then the values are also equal.
Composing each term pₙ in a formal multilinear series with (u, ..., u) where u is a fixed
continuous linear map, gives a new formal multilinear series p.compContinuousLinearMap u.
Instances For
Reinterpret a formal 𝕜'-multilinear series as a formal 𝕜-multilinear series.
Equations
Instances For
Forgetting the zeroth term in a formal multilinear series, and interpreting the following terms
as multilinear maps into E →L[𝕜] F. If p is the Taylor series (HasFTaylorSeriesUpTo) of a
function, then p.shift is the Taylor series of the derivative of the function. Note that the
p.sum of a Taylor series p does not give the original function; for a formal multilinear
series that sums to the derivative of p.sum, see HasFPowerSeriesOnBall.fderiv.
Equations
- p.shift n = (p n.succ).curryRight
Instances For
Adding a zeroth term to a formal multilinear series taking values in E →L[𝕜] F. This
corresponds to starting from a Taylor series (HasFTaylorSeriesUpTo) for the derivative of a
function, and building a Taylor series for the function itself.
Equations
- q.unshift z x = match x with | 0 => (continuousMultilinearCurryFin0 𝕜 E F).symm z | n.succ => (continuousMultilinearCurryRightEquiv' 𝕜 n E F) (q n)
Instances For
Composing each term pₙ in a formal multilinear series with a continuous linear map f on the
left gives a new formal multilinear series f.compFormalMultilinearSeries p whose general term
is f ∘ pₙ.
Equations
- f.compFormalMultilinearSeries p n = f.compContinuousMultilinearMap (p n)
Instances For
Realize a ContinuousMultilinearMap on ∀ i : ι, E i as the evaluation of a
FormalMultilinearSeries by choosing an arbitrary identification ι ≃ Fin (Fintype.card ι).
Equations
- One or more equations did not get rendered due to their size.
Instances For
The index of the first non-zero coefficient in p (or 0 if all coefficients are zero). This
is the order of the isolated zero of an analytic function f at a point if p is the Taylor
series of f at that point.
Instances For
The nth coefficient of p when seen as a power series.
Equations
- p.coeff n = (p n) 1
Instances For
The formal counterpart of dslope, corresponding to the expansion of (f z - f 0) / z. If f
has p as a power series, then dslope f has fslope p as a power series.
Instances For
The formal multilinear series where all terms of positive degree are equal to zero, and the term
of degree zero is c. It is the power series expansion of the constant function equal to c
everywhere.
Equations
- constFormalMultilinearSeries 𝕜 E c x = match x with | 0 => ContinuousMultilinearMap.curry0 𝕜 E c | x => 0
Instances For
Formal power series of a continuous linear map f : E →L[𝕜] F at x : E:
f y = f x + f (y - x).
Equations
- f.fpowerSeries x✝ x = match x with | 0 => ContinuousMultilinearMap.curry0 𝕜 E (f x✝) | 1 => (continuousMultilinearCurryFin1 𝕜 E F).symm f | x => 0