Multivariate polynomials #
This file defines functions for evaluating multivariate polynomials. These include generically evaluating a polynomial given a valuation of all its variables, and more advanced evaluations that allow one to map the coefficients to different rings.
Notation #
In the definitions below, we use the following notation:
σ : Type*
(indexing the variables)R : Type*
[CommSemiring R]
(the coefficients)s : σ →₀ ℕ
, a function fromσ
toℕ
which is zero away from a finite set. This will give rise to a monomial inMvPolynomial σ R
which mathematicians might callX^s
a : R
i : σ
, with corresponding monomialX i
, often denotedX_i
by mathematiciansp : MvPolynomial σ R
Definitions #
eval₂ (f : R → S₁) (g : σ → S₁) p
: given a semiring homomorphism fromR
to another semiringS₁
, and a mapσ → S₁
, evaluatesp
at this valuation, returning a term of typeS₁
. Note thateval₂
can be made usingeval
andmap
(see below), and it has been suggested that sticking toeval
andmap
might make the code less brittle.eval (g : σ → R) p
: given a mapσ → R
, evaluatesp
at this valuation, returning a term of typeR
map (f : R → S₁) p
: returns the multivariate polynomial obtained fromp
by the change of coefficient semiring corresponding tof
aeval (g : σ → S₁) p
: evaluates the multivariate polynomial obtained fromp
by the change of coefficient semiring corresponding tog
(a
stands forAlgebra
)
Evaluate a polynomial p
given a valuation g
of all the variables
and a ring hom f
from the scalar ring to the target
Equations
- MvPolynomial.eval₂ f g p = Finsupp.sum p fun (s : σ →₀ ℕ) (a : R) => f a * s.prod fun (n : σ) (e : ℕ) => g n ^ e
Instances For
MvPolynomial.eval₂
as a RingHom
.
Equations
- MvPolynomial.eval₂Hom f g = { toFun := MvPolynomial.eval₂ f g, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Evaluate a polynomial p
given a valuation f
of all the variables
Equations
Instances For
map f p
maps a polynomial p
across a ring hom f
Equations
Instances For
If f
is a left-inverse of g
then map f
is a left-inverse of map g
.
If f
is a right-inverse of g
then map f
is a right-inverse of map g
.
If f : S₁ →ₐ[R] S₂
is a morphism of R
-algebras, then so is MvPolynomial.map f
.
Equations
- MvPolynomial.mapAlgHom f = { toRingHom := MvPolynomial.map ↑f, commutes' := ⋯ }
Instances For
The algebra of multivariate polynomials #
A map σ → S₁
where S₁
is an algebra over R
generates an R
-algebra homomorphism
from multivariate polynomials over σ
to S₁
.
Equations
- MvPolynomial.aeval f = { toRingHom := MvPolynomial.eval₂Hom (algebraMap R S₁) f, commutes' := ⋯ }
Instances For
Version of aeval
for defining algebra homs out of MvPolynomial σ R
over a smaller base ring
than R
.
Equations
- MvPolynomial.aevalTower f X = { toRingHom := MvPolynomial.eval₂Hom (↑f) X, commutes' := ⋯ }