Additive properties of Hahn series #
If Γ
is ordered and R
has zero, then HahnSeries Γ R
consists of formal series over Γ
with
coefficients in R
, whose supports are partially well-ordered. With further structure on R
and
Γ
, we can add further structure on HahnSeries Γ R
. When R
has an addition operation,
HahnSeries Γ R
also has addition by adding coefficients.
Main Definitions #
- If
R
is a (commutative) additive monoid or group, then so isHahnSeries Γ R
.
References #
- [J. van der Hoeven, Operators on Generalized Power Series][van_der_hoeven]
Equations
- HahnSeries.instAdd = { add := fun (x y : HahnSeries Γ R) => { coeff := x.coeff + y.coeff, isPWO_support' := ⋯ } }
Equations
Alias of HahnSeries.coeff_add'
.
Alias of HahnSeries.coeff_add
.
Alias of HahnSeries.coeff_nsmul
.
addOppositeEquiv
is an additive monoid isomorphism between
Hahn series over Γ
with coefficients in the opposite additive monoid Rᵃᵒᵖ
and the additive opposite of Hahn series over Γ
with coefficients R
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
single
as an additive monoid/group homomorphism
Equations
- HahnSeries.single.addMonoidHom a = { toZeroHom := HahnSeries.single a, map_add' := ⋯ }
Instances For
coeff g
as an additive monoid/group homomorphism
Equations
- HahnSeries.coeff.addMonoidHom g = { toFun := fun (f : HahnSeries Γ R) => f.coeff g, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Equations
Equations
- HahnSeries.instNeg = { neg := fun (x : HahnSeries Γ R) => { coeff := fun (a : Γ) => -x.coeff a, isPWO_support' := ⋯ } }
Equations
Alias of HahnSeries.coeff_neg'
.
Alias of HahnSeries.coeff_neg
.
Alias of HahnSeries.coeff_sub'
.
Alias of HahnSeries.coeff_sub
.
Equations
Equations
- HahnSeries.instSMul = { smul := fun (r : R) (x : HahnSeries Γ V) => { coeff := r • x.coeff, isPWO_support' := ⋯ } }
Alias of HahnSeries.coeff_smul
.
Equations
Equations
Equations
single
as a linear map
Equations
- HahnSeries.single.linearMap a = { toFun := (↑(HahnSeries.single.addMonoidHom a)).toFun, map_add' := ⋯, map_smul' := ⋯ }
Instances For
coeff g
as a linear map
Equations
- HahnSeries.coeff.linearMap g = { toFun := (↑(HahnSeries.coeff.addMonoidHom g)).toFun, map_add' := ⋯, map_smul' := ⋯ }
Instances For
Extending the domain of Hahn series is a linear map.
Equations
- HahnSeries.embDomainLinearMap f = { toFun := HahnSeries.embDomain f, map_add' := ⋯, map_smul' := ⋯ }