Evaluation of multivariate power series #
Let σ
, R
, S
be types, with CommRing R
, CommRing S
.
One assumes that IsTopologicalRing R
and UniformAddGroup R
,
and that S
is a complete and separated topological R
-algebra,
with IsLinearTopology S S
, which means there is a basis of neighborhoods of 0
consisting of ideals.
Given φ : R →+* S
, a : σ → S
, and f : MvPowerSeries σ R
,
MvPowerSeries.eval₂ f φ a
is the evaluation of the multivariate power series f
at a
.
It f
is (the coercion of) a polynomial, it coincides with the evaluation of that polynomial.
Otherwise, it is defined by density from polynomials;
its values are irrelevant unless φ
is continuous and a
satisfies two conditions
bundled in MvPowerSeries.HasEval a
:
- for all
s : σ
,a s
is topologically nilpotent, meaning that(a s) ^ n
tends to 0 whenn
tends to infinity - when
a s
tends to zero for the filter of cofinite subsets ofσ
.
Under Continuous φ
and HasEval a
, the following lemmas furnish the properties of evaluation:
MvPowerSeries.eval₂Hom
: the evaluation of multivariate power series, as a ring morphism,MvPowerSeries.aeval
: the evaluation map as an algebra morphismMvPowerSeries.uniformContinuous_eval₂
: uniform continuity of the evaluationMvPowerSeries.continuous_eval₂
: continuity of the evaluationMvPowerSeries.eval₂_eq_tsum
: the evaluation is given by the sum of its monomials, evaluated.
Families at which power series can be consistently evaluated
- hpow (s : σ) : IsTopologicallyNilpotent (a s)
- tendsto_zero : Filter.Tendsto a Filter.cofinite (nhds 0)
Instances For
The domain of evaluation of MvPowerSeries
, as an ideal
Equations
- MvPowerSeries.hasEvalIdeal = { carrier := {a : σ → S | MvPowerSeries.HasEval a}, add_mem' := ⋯, zero_mem' := ⋯, smul_mem' := ⋯ }
Instances For
The inclusion of polynomials into power series has dense image
Evaluation of a multivariate power series at f
at a point a : σ → S
.
It coincides with the evaluation of f
as a polynomial if f
is the coercion of a polynomial.
Otherwise, it is only relevant if φ
is continuous and HasEval a
.
Equations
- MvPowerSeries.eval₂ φ a f = if H : ∃ (p : MvPolynomial σ R), ↑p = f then MvPolynomial.eval₂ φ a H.choose else ⋯.extend (MvPolynomial.eval₂ φ a) f
Instances For
Evaluation of power series at adequate elements, as a RingHom
Equations
- MvPowerSeries.eval₂Hom hφ ha = IsDenseInducing.extendRingHom ⋯ ⋯ ⋯
Instances For
Evaluation of power series at adequate elements, as an AlgHom
Equations
- MvPowerSeries.aeval ha = { toRingHom := MvPowerSeries.eval₂Hom ⋯ ha, commutes' := ⋯ }