Documentation

Mathlib.RingTheory.HahnSeries.Addition

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 #

References #

instance HahnSeries.instAdd {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] :
Equations
  • HahnSeries.instAdd = { add := fun (x y : HahnSeries Γ R) => { coeff := x.coeff + y.coeff, isPWO_support' := } }
instance HahnSeries.instAddMonoid {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] :
Equations
@[simp]
theorem HahnSeries.add_coeff' {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] {x y : HahnSeries Γ R} :
(x + y).coeff = x.coeff + y.coeff
theorem HahnSeries.add_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] {x y : HahnSeries Γ R} {a : Γ} :
(x + y).coeff a = x.coeff a + y.coeff a
@[simp]
theorem HahnSeries.nsmul_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] {x : HahnSeries Γ R} {n : } :
(n x).coeff = n x.coeff
@[simp]
theorem HahnSeries.map_add {Γ : Type u_1} {R : Type u_3} {S : Type u_4} [PartialOrder Γ] [AddMonoid R] [AddMonoid S] (f : R →+ S) {x y : HahnSeries Γ R} :
(x + y).map f = x.map f + y.map f

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
    theorem HahnSeries.addOppositeEquiv_symm_apply_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : (HahnSeries Γ R)ᵃᵒᵖ) (a : Γ) :
    (HahnSeries.addOppositeEquiv.symm x).coeff a = AddOpposite.op ((AddOpposite.unop x).coeff a)
    theorem HahnSeries.addOppositeEquiv_apply {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : HahnSeries Γ Rᵃᵒᵖ) :
    HahnSeries.addOppositeEquiv x = AddOpposite.op { coeff := fun (a : Γ) => AddOpposite.unop (x.coeff a), isPWO_support' := }
    @[simp]
    theorem HahnSeries.addOppositeEquiv_support {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : HahnSeries Γ Rᵃᵒᵖ) :
    (AddOpposite.unop (HahnSeries.addOppositeEquiv x)).support = x.support
    @[simp]
    theorem HahnSeries.addOppositeEquiv_symm_support {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : (HahnSeries Γ R)ᵃᵒᵖ) :
    (HahnSeries.addOppositeEquiv.symm x).support = (AddOpposite.unop x).support
    @[simp]
    theorem HahnSeries.addOppositeEquiv_orderTop {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : HahnSeries Γ Rᵃᵒᵖ) :
    (AddOpposite.unop (HahnSeries.addOppositeEquiv x)).orderTop = x.orderTop
    @[simp]
    theorem HahnSeries.addOppositeEquiv_symm_orderTop {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : (HahnSeries Γ R)ᵃᵒᵖ) :
    (HahnSeries.addOppositeEquiv.symm x).orderTop = (AddOpposite.unop x).orderTop
    @[simp]
    theorem HahnSeries.addOppositeEquiv_leadingCoeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : HahnSeries Γ Rᵃᵒᵖ) :
    (AddOpposite.unop (HahnSeries.addOppositeEquiv x)).leadingCoeff = AddOpposite.unop x.leadingCoeff
    @[simp]
    theorem HahnSeries.addOppositeEquiv_symm_leadingCoeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (x : (HahnSeries Γ R)ᵃᵒᵖ) :
    (HahnSeries.addOppositeEquiv.symm x).leadingCoeff = AddOpposite.op (AddOpposite.unop x).leadingCoeff
    theorem HahnSeries.support_add_subset {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] {x y : HahnSeries Γ R} :
    (x + y).support x.support y.support
    theorem HahnSeries.min_le_min_add {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hx : x 0) (hy : y 0) (hxy : x + y 0) :
    .min .min .min
    theorem HahnSeries.min_orderTop_le_orderTop_add {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} :
    x.orderTop y.orderTop (x + y).orderTop
    theorem HahnSeries.min_order_le_order_add {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [Zero Γ] [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : x + y 0) :
    x.order y.order (x + y).order
    theorem HahnSeries.orderTop_add_eq_left {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : x.orderTop < y.orderTop) :
    (x + y).orderTop = x.orderTop
    theorem HahnSeries.orderTop_add_eq_right {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : y.orderTop < x.orderTop) :
    (x + y).orderTop = y.orderTop
    theorem HahnSeries.leadingCoeff_add_eq_left {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : x.orderTop < y.orderTop) :
    (x + y).leadingCoeff = x.leadingCoeff
    theorem HahnSeries.leadingCoeff_add_eq_right {R : Type u_3} [AddMonoid R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : y.orderTop < x.orderTop) :
    (x + y).leadingCoeff = y.leadingCoeff
    def HahnSeries.single.addMonoidHom {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (a : Γ) :

    single as an additive monoid/group homomorphism

    Equations
    Instances For
      @[simp]
      theorem HahnSeries.single.addMonoidHom_apply_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (a : Γ) (r : R) (j : Γ) :
      def HahnSeries.coeff.addMonoidHom {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (g : Γ) :

      coeff g as an additive monoid/group homomorphism

      Equations
      Instances For
        @[simp]
        theorem HahnSeries.coeff.addMonoidHom_apply {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] (g : Γ) (f : HahnSeries Γ R) :
        theorem HahnSeries.embDomain_add {Γ : Type u_1} {Γ' : Type u_2} {R : Type u_3} [PartialOrder Γ] [AddMonoid R] [PartialOrder Γ'] (f : Γ ↪o Γ') (x y : HahnSeries Γ R) :
        Equations
        instance HahnSeries.instNeg {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] :
        Equations
        • HahnSeries.instNeg = { neg := fun (x : HahnSeries Γ R) => { coeff := fun (a : Γ) => -x.coeff a, isPWO_support' := } }
        instance HahnSeries.instAddGroup {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] :
        Equations
        @[simp]
        theorem HahnSeries.neg_coeff' {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x : HahnSeries Γ R} :
        (-x).coeff = -x.coeff
        theorem HahnSeries.neg_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x : HahnSeries Γ R} {a : Γ} :
        (-x).coeff a = -x.coeff a
        @[simp]
        theorem HahnSeries.support_neg {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x : HahnSeries Γ R} :
        (-x).support = x.support
        @[simp]
        theorem HahnSeries.map_neg {Γ : Type u_1} {R : Type u_3} {S : Type u_4} [PartialOrder Γ] [AddGroup R] [AddGroup S] (f : R →+ S) {x : HahnSeries Γ R} :
        (-x).map f = -x.map f
        theorem HahnSeries.orderTop_neg {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x : HahnSeries Γ R} :
        (-x).orderTop = x.orderTop
        @[simp]
        theorem HahnSeries.order_neg {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] [Zero Γ] {f : HahnSeries Γ R} :
        (-f).order = f.order
        @[simp]
        theorem HahnSeries.sub_coeff' {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x y : HahnSeries Γ R} :
        (x - y).coeff = x.coeff - y.coeff
        theorem HahnSeries.sub_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] [AddGroup R] {x y : HahnSeries Γ R} {a : Γ} :
        (x - y).coeff a = x.coeff a - y.coeff a
        @[simp]
        theorem HahnSeries.map_sub {Γ : Type u_1} {R : Type u_3} {S : Type u_4} [PartialOrder Γ] [AddGroup R] [AddGroup S] (f : R →+ S) {x y : HahnSeries Γ R} :
        (x - y).map f = x.map f - y.map f
        theorem HahnSeries.min_orderTop_le_orderTop_sub {R : Type u_3} [AddGroup R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} :
        x.orderTop y.orderTop (x - y).orderTop
        theorem HahnSeries.orderTop_sub {R : Type u_3} [AddGroup R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : x.orderTop < y.orderTop) :
        (x - y).orderTop = x.orderTop
        theorem HahnSeries.leadingCoeff_sub {R : Type u_3} [AddGroup R] {Γ : Type u_7} [LinearOrder Γ] {x y : HahnSeries Γ R} (hxy : x.orderTop < y.orderTop) :
        (x - y).leadingCoeff = x.leadingCoeff
        Equations
        instance HahnSeries.instSMul {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Zero V] [SMulZeroClass R V] :
        SMul R (HahnSeries Γ V)
        Equations
        • HahnSeries.instSMul = { smul := fun (r : R) (x : HahnSeries Γ V) => { coeff := r x.coeff, isPWO_support' := } }
        @[simp]
        theorem HahnSeries.smul_coeff {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Zero V] [SMulZeroClass R V] {r : R} {x : HahnSeries Γ V} {a : Γ} :
        (r x).coeff a = r x.coeff a
        instance HahnSeries.instSMulZeroClass {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Zero V] [SMulZeroClass R V] :
        Equations
        theorem HahnSeries.orderTop_smul_not_lt {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Zero V] [SMulZeroClass R V] (r : R) (x : HahnSeries Γ V) :
        ¬(r x).orderTop < x.orderTop
        theorem HahnSeries.order_smul_not_lt {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Zero V] [SMulZeroClass R V] [Zero Γ] (r : R) (x : HahnSeries Γ V) (h : r x 0) :
        ¬(r x).order < x.order
        theorem HahnSeries.le_order_smul {R : Type u_3} {V : Type u_7} [Zero V] [SMulZeroClass R V] {Γ : Type u_8} [Zero Γ] [LinearOrder Γ] (r : R) (x : HahnSeries Γ V) (h : r x 0) :
        x.order (r x).order
        instance HahnSeries.instDistribMulAction {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Monoid R] [AddMonoid V] [DistribMulAction R V] :
        Equations
        instance HahnSeries.instIsScalarTower {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Monoid R] [AddMonoid V] [DistribMulAction R V] {S : Type u_8} [Monoid S] [DistribMulAction S V] [SMul R S] [IsScalarTower R S V] :
        Equations
        • =
        instance HahnSeries.instSMulCommClass {Γ : Type u_1} {R : Type u_3} [PartialOrder Γ] {V : Type u_7} [Monoid R] [AddMonoid V] [DistribMulAction R V] {S : Type u_8} [Monoid S] [DistribMulAction S V] [SMulCommClass R S V] :
        Equations
        • =
        instance HahnSeries.instModule {Γ : Type u_1} {R : Type u_3} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] :
        Equations
        def HahnSeries.single.linearMap {Γ : Type u_1} {R : Type u_3} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] (a : Γ) :

        single as a linear map

        Equations
        Instances For
          @[simp]
          theorem HahnSeries.single.linearMap_apply {Γ : Type u_1} {R : Type u_3} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] (a : Γ) (a✝ : V) :
          def HahnSeries.coeff.linearMap {Γ : Type u_1} {R : Type u_3} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] (g : Γ) :

          coeff g as a linear map

          Equations
          Instances For
            @[simp]
            theorem HahnSeries.coeff.linearMap_apply {Γ : Type u_1} {R : Type u_3} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] (g : Γ) (a✝ : HahnSeries Γ V) :
            @[simp]
            theorem HahnSeries.map_smul {Γ : Type u_1} {R : Type u_3} {U : Type u_5} {V : Type u_6} [PartialOrder Γ] [Semiring R] [AddCommMonoid V] [Module R V] [AddCommMonoid U] [Module R U] (f : U →ₗ[R] V) {r : R} {x : HahnSeries Γ U} :
            (r x).map f = r x.map f
            theorem HahnSeries.embDomain_smul {Γ : Type u_1} {Γ' : Type u_2} {R : Type u_3} [PartialOrder Γ] [Semiring R] [PartialOrder Γ'] (f : Γ ↪o Γ') (r : R) (x : HahnSeries Γ R) :
            def HahnSeries.embDomainLinearMap {Γ : Type u_1} {Γ' : Type u_2} {R : Type u_3} [PartialOrder Γ] [Semiring R] [PartialOrder Γ'] (f : Γ ↪o Γ') :

            Extending the domain of Hahn series is a linear map.

            Equations
            Instances For
              @[simp]
              theorem HahnSeries.embDomainLinearMap_apply {Γ : Type u_1} {Γ' : Type u_2} {R : Type u_3} [PartialOrder Γ] [Semiring R] [PartialOrder Γ'] (f : Γ ↪o Γ') (a✝ : HahnSeries Γ R) :