Documentation

Mathlib.RingTheory.MvPowerSeries.Order

Order of multivariate power series #

We work with MvPowerSeries σ R, for Semiring R, and w : σ → ℕ

Weighted Order #

Order #

Homogeneous components #

NOTE: Under Finite σ, one can use Finsupp.finite_of_degree_le and Finsupp.finite_of_weight_le to show that they have finite support, hence correspond to MvPolynomial.

However, when σ is finite, this is not necessarily an MvPolynomial. (For example: the homogeneous components of degree 1 of the multivariate power series, all of which coefficients are 1, is the sum of all indeterminates.)

TODO: Define a coercion to MvPolynomial.

theorem MvPowerSeries.ne_zero_iff_exists_coeff_ne_zero_and_weight {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} :
f 0 ∃ (n : ) (d : σ →₀ ), (MvPowerSeries.coeff R d) f 0 (Finsupp.weight w) d = n
def MvPowerSeries.weightedOrder {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) (f : MvPowerSeries σ R) :

The weighted order of a mv_power_series

Equations
Instances For
    @[simp]
    theorem MvPowerSeries.weightedOrder_zero {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) :
    @[simp]
    theorem MvPowerSeries.weightedOrder_eq_top_iff {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} :

    The 0 power series is the unique power series with infinite order.

    If the order of a formal power series f is finite, then some coefficient of weight equal to the order of f is nonzero.

    theorem MvPowerSeries.weightedOrder_le {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} {d : σ →₀ } (h : (MvPowerSeries.coeff R d) f 0) :

    If the dth coefficient of a formal power series is nonzero, then the weighted order of the power series is less than or equal to weight d w.

    theorem MvPowerSeries.coeff_eq_zero_of_lt_weightedOrder {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} {d : σ →₀ } (h : ((Finsupp.weight w) d) < MvPowerSeries.weightedOrder w f) :

    The nth coefficient of a formal power series is 0 if n is strictly smaller than the order of the power series.

    theorem MvPowerSeries.nat_le_weightedOrder {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} {n : } (h : ∀ (d : σ →₀ ), (Finsupp.weight w) d < n(MvPowerSeries.coeff R d) f = 0) :

    The order of a formal power series is at least n if the dth coefficient is 0 for all d such that weight w d < n.

    theorem MvPowerSeries.le_weightedOrder {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} {n : ℕ∞} (h : ∀ (d : σ →₀ ), ((Finsupp.weight w) d) < n(MvPowerSeries.coeff R d) f = 0) :

    The order of a formal power series is at least n if the dth coefficient is 0 for all d such that weight w d < n.

    theorem MvPowerSeries.weightedOrder_eq_nat {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {f : MvPowerSeries σ R} {n : } :
    MvPowerSeries.weightedOrder w f = n (∃ (d : σ →₀ ), (MvPowerSeries.coeff R d) f 0 (Finsupp.weight w) d = n) ∀ (d : σ →₀ ), (Finsupp.weight w) d < n(MvPowerSeries.coeff R d) f = 0

    The order of a formal power series is exactly n if and only if some coefficient of weight n is nonzero, and the dth coefficient is 0 for all d such that weight w d < n.

    theorem MvPowerSeries.weightedOrder_monomial {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {d : σ →₀ } {a : R} [Decidable (a = 0)] :

    The weighted_order of the monomial a*X^d is infinite if a = 0 and weight w d otherwise.

    theorem MvPowerSeries.weightedOrder_monomial_of_ne_zero {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) {d : σ →₀ } {a : R} (h : a 0) :

    The order of the monomial a*X^n is n if a ≠ 0.

    The order of the sum of two formal power series is at least the minimum of their orders.

    The weighted_order of the sum of two formal power series is the minimum of their orders if their orders differ.

    The weighted_order of the product of two formal power series is at least the sum of their orders.

    Alias of MvPowerSeries.le_weightedOrder_mul.


    The weighted_order of the product of two formal power series is at least the sum of their orders.

    theorem MvPowerSeries.coeff_mul_left_one_sub_of_lt_weightedOrder {σ : Type u_1} (w : σ) {R : Type u_3} [Ring R] {f g : MvPowerSeries σ R} {d : σ →₀ } (h : ((Finsupp.weight w) d) < MvPowerSeries.weightedOrder w g) :
    (MvPowerSeries.coeff R d) (f * (1 - g)) = (MvPowerSeries.coeff R d) f
    theorem MvPowerSeries.coeff_mul_right_one_sub_of_lt_weightedOrder {σ : Type u_1} (w : σ) {R : Type u_3} [Ring R] {f g : MvPowerSeries σ R} {d : σ →₀ } (h : ((Finsupp.weight w) d) < MvPowerSeries.weightedOrder w g) :
    (MvPowerSeries.coeff R d) ((1 - g) * f) = (MvPowerSeries.coeff R d) f
    theorem MvPowerSeries.coeff_mul_prod_one_sub_of_lt_weightedOrder {σ : Type u_1} (w : σ) {R : Type u_4} {ι : Type u_5} [CommRing R] (d : σ →₀ ) (s : Finset ι) (f : MvPowerSeries σ R) (g : ιMvPowerSeries σ R) :
    (∀ is, ((Finsupp.weight w) d) < MvPowerSeries.weightedOrder w (g i))(MvPowerSeries.coeff R d) (f * is, (1 - g i)) = (MvPowerSeries.coeff R d) f
    theorem MvPowerSeries.eq_zero_iff_forall_coeff_eq_zero_and {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} :
    f = 0 ∀ (d : σ →₀ ), (MvPowerSeries.coeff R d) f = 0
    theorem MvPowerSeries.ne_zero_iff_exists_coeff_ne_zero_and_degree {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} :
    f 0 ∃ (n : ) (d : σ →₀ ), (MvPowerSeries.coeff R d) f 0 d.degree = n
    def MvPowerSeries.order {σ : Type u_1} {R : Type u_2} [Semiring R] (f : MvPowerSeries σ R) :

    The order of a mv_power_series

    Equations
    Instances For
      @[simp]
      theorem MvPowerSeries.order_zero {σ : Type u_1} {R : Type u_2} [Semiring R] :
      theorem MvPowerSeries.ne_zero_iff_order_finite {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} :
      f 0 f.order.toNat = f.order
      @[simp]
      theorem MvPowerSeries.order_eq_top_iff {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} :
      f.order = f = 0

      The 0 power series is the unique power series with infinite order.

      theorem MvPowerSeries.exists_coeff_ne_zero_and_order {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} (h : f.order.toNat = f.order) :
      ∃ (d : σ →₀ ), (MvPowerSeries.coeff R d) f 0 d.degree = f.order

      If the order of a formal power series f is finite, then some coefficient of degree the order of f is nonzero.

      theorem MvPowerSeries.order_le {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} {d : σ →₀ } (h : (MvPowerSeries.coeff R d) f 0) :
      f.order d.degree

      If the dth coefficient of a formal power series is nonzero, then the order of the power series is less than or equal to degree d.

      theorem MvPowerSeries.coeff_of_lt_order {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} {d : σ →₀ } (h : d.degree < f.order) :

      The nth coefficient of a formal power series is 0 if n is strictly smaller than the order of the power series.

      theorem MvPowerSeries.nat_le_order {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} {n : } (h : ∀ (d : σ →₀ ), d.degree < n(MvPowerSeries.coeff R d) f = 0) :
      n f.order

      The order of a formal power series is at least n if the dth coefficient is 0 for all d such that degree d < n.

      theorem MvPowerSeries.le_order {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} {n : ℕ∞} (h : ∀ (d : σ →₀ ), d.degree < n(MvPowerSeries.coeff R d) f = 0) :
      n f.order

      The order of a formal power series is at least n if the dth coefficient is 0 for all d such that degree d < n.

      theorem MvPowerSeries.order_eq_nat {σ : Type u_1} {R : Type u_2} [Semiring R] {f : MvPowerSeries σ R} {n : } :
      f.order = n (∃ (d : σ →₀ ), (MvPowerSeries.coeff R d) f 0 d.degree = n) ∀ (d : σ →₀ ), d.degree < n(MvPowerSeries.coeff R d) f = 0

      The order of a formal power series is exactly n some coefficient of degree n is nonzero, and the dth coefficient is 0 for all d such that degree d < n.

      theorem MvPowerSeries.order_monomial {σ : Type u_1} {R : Type u_2} [Semiring R] {d : σ →₀ } {a : R} [Decidable (a = 0)] :
      ((MvPowerSeries.monomial R d) a).order = if a = 0 then else d.degree

      The order of the monomial a*X^d is infinite if a = 0 and degree d otherwise.

      theorem MvPowerSeries.order_monomial_of_ne_zero {σ : Type u_1} {R : Type u_2} [Semiring R] {d : σ →₀ } {a : R} (h : a 0) :
      ((MvPowerSeries.monomial R d) a).order = d.degree

      The order of the monomial a*X^n is n if a ≠ 0.

      theorem MvPowerSeries.min_order_le_add {σ : Type u_1} {R : Type u_2} [Semiring R] {f g : MvPowerSeries σ R} :
      f.order g.order (f + g).order

      The order of the sum of two formal power series is at least the minimum of their orders.

      theorem MvPowerSeries.order_add_of_order_ne {σ : Type u_1} {R : Type u_2} [Semiring R] {f g : MvPowerSeries σ R} (h : f.order g.order) :
      (f + g).order = f.order g.order

      The order of the sum of two formal power series is the minimum of their orders if their orders differ.

      theorem MvPowerSeries.le_order_mul {σ : Type u_1} {R : Type u_2} [Semiring R] {f g : MvPowerSeries σ R} :
      f.order + g.order (f * g).order

      The order of the product of two formal power series is at least the sum of their orders.

      theorem MvPowerSeries.order_mul_ge {σ : Type u_1} {R : Type u_2} [Semiring R] {f g : MvPowerSeries σ R} :
      f.order + g.order (f * g).order

      Alias of MvPowerSeries.le_order_mul.


      The order of the product of two formal power series is at least the sum of their orders.

      theorem MvPowerSeries.coeff_mul_left_one_sub_of_lt_order {σ : Type u_1} {R : Type u_3} [Ring R] {f g : MvPowerSeries σ R} (d : σ →₀ ) (h : d.degree < g.order) :
      (MvPowerSeries.coeff R d) (f * (1 - g)) = (MvPowerSeries.coeff R d) f
      theorem MvPowerSeries.coeff_mul_right_one_sub_of_lt_order {σ : Type u_1} {R : Type u_3} [Ring R] {f g : MvPowerSeries σ R} (d : σ →₀ ) (h : d.degree < g.order) :
      (MvPowerSeries.coeff R d) ((1 - g) * f) = (MvPowerSeries.coeff R d) f
      theorem MvPowerSeries.coeff_mul_prod_one_sub_of_lt_order {σ : Type u_1} {R : Type u_4} {ι : Type u_5} [CommRing R] (d : σ →₀ ) (s : Finset ι) (f : MvPowerSeries σ R) (g : ιMvPowerSeries σ R) :
      (∀ is, d.degree < (g i).order)(MvPowerSeries.coeff R d) (f * is, (1 - g i)) = (MvPowerSeries.coeff R d) f
      def MvPowerSeries.weightedHomogeneousComponent {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) (p : ) :

      The weighted homogeneous components of an MvPowerSeries f.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        theorem MvPowerSeries.coeff_weightedHomogeneousComponent {σ : Type u_1} {R : Type u_2} [Semiring R] (w : σ) (p : ) (d : σ →₀ ) (f : MvPowerSeries σ R) :
        theorem MvPowerSeries.coeff_homogeneousComponent {σ : Type u_1} {R : Type u_2} [Semiring R] (p : ) (d : σ →₀ ) (f : MvPowerSeries σ R) :
        (MvPowerSeries.coeff R d) ((MvPowerSeries.homogeneousComponent p) f) = if d.degree = p then (MvPowerSeries.coeff R d) f else 0