Division algorithm with respect to monomial orders #
We provide a division algorithm with respect to monomial orders in polynomial rings.
Let R
be a commutative ring, σ
a type of indeterminates and m : MonomialOrder σ
a monomial ordering on σ →₀ ℕ
.
Consider a family of polynomials b : ι → MvPolynomial σ R
with invertible leading coefficients
(with respect to m
) : we assume hb : ∀ i, IsUnit (m.lCoeff (b i))
).
MonomialOrder.div hb f
furnishes- a finitely supported family
g : ι →₀ MvPolynomial σ R
- and a “remainder”
r : MvPolynomial σ R
such that the three properties hold: (1) One hasf = ∑ (g i) * (b i) + r
(2) For everyi
,m.degree ((g i) * (b i)
is less than or equal to that off
(3) For everyi
, every monomial in the support ofr
is strictly smaller than the leading term ofb i
,
- a finitely supported family
The proof is done by induction, using two standard constructions
MonomialOrder.subLTerm f
deletes the leading term of a polynomialf
MonomialOrder.reduce hb f
subtracts fromf
the appropriate multiple ofb : MvPolynomial σ R
, providedIsUnit (m.lCoeff b)
.MonomialOrder.div_set
is the variant ofMonomialOrder.div
for a set of polynomials.
Reference : [BW93] #
TODO #
Delete the leading term in a multivariate polynomial (for some monomial order)
Equations
- m.subLTerm f = f - (MvPolynomial.monomial (m.degree f)) (m.lCoeff f)
Instances For
Reduce a polynomial modulo a polynomial with unit leading term (for some monomial order)