Order of multivariate power series #
We work with MvPowerSeries σ R, for Semiring R, and w : σ → ℕ
Weighted Order #
MvPowerSeries.weightedOrder: the weighted order of a multivariate power series, with respect tow, as an element ofℕ∞.MvPowerSeries.weightedOrder_zero: the weighted order of0is0.MvPowerSeries.ne_zero_iff_weightedOrder_finite: a multivariate power series is nonzero if and only if its weighted order is finite.MvPowerSeries.exists_coeff_ne_zero_of_weightedOrder: if the weighted order is finite, then there exists a nonzero coefficient of weight the weighted order.MvPowerSeries.weightedOrder_le: if a coefficient is nonzero, then the weighted order is at most the weight of that exponent.MvPowerSeries.coeff_eq_zero_of_lt_weightedOrder: all coefficients of weights strictly less than the weighted order vanishMvPowerSeries.weightedOrder_eq_top_iff: the weighted order offis⊤if and onlyf = 0.MvPowerSeries.nat_le_weightedOrder: if all coefficients of weight< nvanish, then the weighted order is at leastn.MvPowerSeries.weightedOrder_eq_nat_iff: the weighted order is some integerniff there exists a nonzero coefficient of weightn, and all coefficients of strictly smaller weight vanish.MvPowerSeries.weightedOrder_monomial,MvPowerSeries.weightedOrder_monomial_of_ne_zero: the weighted order of a monomial, of a monomial with nonzero coefficient.MvPowerSeries.min_weightedOrder_le_add: the order of the sum of two multivariate power series is at least the minimum of their orders.MvPowerSeries.weightedOrder_add_of_weightedOrder_ne: the weighted_order of the sum of two formal power series is the minimum of their orders if their orders differ.MvPowerSeries.le_weightedOrder_mul: the weighted_order of the product of two formal power series is at least the sum of their orders.MvPowerSeries.coeff_mul_left_one_sub_of_lt_weightedOrder,MvPowerSeries.coeff_mul_right_one_sub_of_lt_weightedOrder: the coefficients off * (1 - g)and(1 - g) * fin weights strictly less than the weighted order ofg.MvPowerSeries.coeff_mul_prod_one_sub_of_lt_weightedOrder: the coefficients off * Π i in s, (1 - g i), in weights strictly less than the weighted orders ofg i, fori ∈ s.
Order #
MvPowerSeries.order: the weighted order, forw = (1 : σ → ℕ).MvPowerSeries.ne_zero_iff_order_finite:fis nonzero iff its order is finite.MvPowerSeries.order_eq_top_iff: the order offis infinite ifff = 0.MvPowerSeries.exists_coeff_ne_zero_of_order: if the order is finite, then there exists a nonzero coefficient of degree equal to the order.MvPowerSeries.order_le: if a coefficient of some degree is nonzero, then the order is at least that degree.MvPowerSeries.nat_le_order: if all coefficients of degree strictly smaller than some integer vanish, then the order is at least that integer.MvPowerSeries.order_eq_nat_iff: the order of a power series is an integerniff there exists a nonzero coefficient in that degree, and all coefficients below that degree vanish.MvPowerSeries.order_monomial,MvPowerSeries.order_monomial_of_ne_zero: the order of a monomial, with a nonzero coefficientMvPowerSeries.min_order_le_add: the order of a sum of two power series is at least the minimum of their orders.MvPowerSeries.order_add_of_order_ne: the order of a sum of two power series of distinct orders is the minimum of their orders.MvPowerSeries.order_mul_ge: the order of a product of two power series is at least the sum of their orders.MvPowerSeries.coeff_mul_left_one_sub_of_lt_order,MvPowerSeries.coeff_mul_right_one_sub_of_lt_order: the coefficients off * (1 - g)and(1 - g) * fbelow the order ofgcoincide with that off.MvPowerSeries.coeff_mul_prod_one_sub_of_lt_order: the coefficients off * Π i in s, (1 - g i)coincide with that offbelow the minimum of the orders of theg i, fori ∈ s.
Homogeneous components #
MvPowerSeries.weightedHomogeneousComponent,MvPowerSeries.homogeneousComponent: the power series which is the sum of all monomials of given weighted degree, resp. degree.
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.
The weighted order of a mv_power_series
Equations
- MvPowerSeries.weightedOrder w f = if x : f = 0 then ⊤ else ↑(Nat.find ⋯)
Instances For
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.
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.
The nth coefficient of a formal power series is 0 if n is strictly
smaller than the order of the power series.
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.
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.
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.
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.
The order of a mv_power_series
Equations
- f.order = MvPowerSeries.weightedOrder (fun (x : σ) => 1) f
Instances For
If the order of a formal power series f is finite,
then some coefficient of degree the order of f is nonzero.
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.
Alias of MvPowerSeries.le_order_mul.
The order of the product of two formal power series is at least the sum of their orders.
Weighted homogeneous power series
Equations
- MvPowerSeries.IsWeightedHomogeneous w f p = ∀ {d : σ →₀ ℕ}, (MvPowerSeries.coeff d) f ≠ 0 → (Finsupp.weight w) d = p
Instances For
The weighted homogeneous components of an MvPowerSeries f.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Homogeneous power series
Equations
Instances For
The homogeneous components of an MvPowerSeries