Polynomials with degree strictly less than n
#
This file contains the properties of the submodule of polynomials of degree less than n
in a
(semi)ring R
, denoted R[X]_n
.
Main definitions/lemmas #
degreeLT.basis R n
: a basis forR[X]_n
the submodule of polynomials with degree< n
, given by the monomialsX^i
fori < n
.degreeLT.basisProd R m n
: a basis forR[X]_m × R[X]_n
, which is the sum of two instances of the basis given above.degreeLT.addLinearEquiv R m n
: an isomorphism betweenR[X]_(m+n)
andR[X]_m × R[X]_n
, given by the fact that the bases are both indexed byFin (m+n)
. This is used for the Sylvester matrix, which is the matrix representing the Sylvester map between these two spaces, in a future file.taylorLinear r n
: The linear automorphism induced bytaylor r
onR[X]_n
which sendsX
toX + r
and preserves degrees.
The R
-submodule of R[X]
consisting of polynomials of degree < n
.
Equations
- Polynomial.«term_[X]__» = Lean.ParserDescr.trailingNode `Polynomial.«term_[X]__» 9000 0 (Lean.ParserDescr.binary `andthen (Lean.ParserDescr.symbol "[X]_") (Lean.ParserDescr.cat `term 1023))
Instances For
Basis for R[X]_n
given by X^i
with i < n
.
Equations
Instances For
Basis for R[X]_m × R[X]_n
.
Equations
Instances For
An isomorphism between R[X]_(m+n)
and R[X]_m × R[X]_n
given by the fact that the bases are
both indexed by Fin (m+n)
.
Equations
- Polynomial.degreeLT.addLinearEquiv R m n = (Polynomial.degreeLT.basis R (m + n)).equiv (Polynomial.degreeLT.basisProd R m n) (Equiv.refl (Fin (m + n)))
Instances For
The map taylor r
induces an automorphism of the module R[X]_n
of polynomials of
degree < n
.
Equations
- Polynomial.taylorLinearEquiv r n = (↑(Polynomial.taylorEquiv r)).ofSubmodules (Polynomial.degreeLT R n) (Polynomial.degreeLT R n) ⋯