Algebraic elements and algebraic extensions #
An element of an R-algebra is algebraic over R if it is the root of a nonzero polynomial. An R-algebra is algebraic over R if and only if all its elements are algebraic over R.
Main definitions #
IsAlgebraic
: algebraic elements of an algebra.Transcendental
: transcendental elements of an algebra are those that are not algebraic.Subalgebra.IsAlgebraic
: a subalgebra is algebraic if all its elements are algebraic.Algebra.IsAlgebraic
: an algebra is algebraic if all its elements are algebraic.Algebra.Transcendental
: an algebra is transcendental if some element is transcendental.
Main results #
transcendental_iff
: an elementx : A
is transcendental overR
iff out ofR[X]
only the zero polynomial evaluates to 0 atx
.Subalgebra.isAlgebraic_iff
: a subalgebra is algebraic iff it is algebraic as an algebra.
An element of an R-algebra is algebraic over R if it is a root of a nonzero polynomial with coefficients in R.
Stacks Tag 09GC (Algebraic elements)
Equations
- IsAlgebraic R x = ∃ (p : Polynomial R), p ≠ 0 ∧ (Polynomial.aeval x) p = 0
Instances For
An element of an R-algebra is transcendental over R if it is not algebraic over R.
Equations
- Transcendental R x = ¬IsAlgebraic R x
Instances For
An element x
is transcendental over R
if and only if for any polynomial p
,
Polynomial.aeval x p = 0
implies p = 0
. This is similar to algebraicIndependent_iff
.
A subalgebra is algebraic if all its elements are algebraic.
Equations
- S.IsAlgebraic = ∀ x ∈ S, IsAlgebraic R x
Instances For
An algebra is algebraic if all its elements are algebraic.
Stacks Tag 09GC (Algebraic extensions)
- isAlgebraic : ∀ (x : A), IsAlgebraic R x
Instances
An algebra is transcendental if some element is transcendental.
- transcendental : ∃ (x : A), Transcendental R x
Instances
A subalgebra is algebraic if and only if it is algebraic as an algebra.
An algebra is algebraic if and only if it is algebraic as a subalgebra.