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. The main result in this file proves transitivity of algebraicity: a tower of algebraic field extensions is algebraic.
If r : A
and f : R[X]
are transcendental over R
, then Polynomial.aeval r f
is also
transcendental over R
. For the converse, see Transcendental.of_aeval
and
transcendental_aeval_iff
.
If Polynomial.aeval r f
is transcendental over R
, then f : R[X]
is also
transcendental over R
. In fact, the r
is also transcendental over R
provided that R
is a field (see transcendental_aeval_iff
).
An element x
is transcendental over R
if and only if the map Polynomial.aeval x
is injective. This is similar to algebraicIndependent_iff_injective_aeval
.
An element x
is transcendental over R
if and only if the kernel of the ring homomorphism
Polynomial.aeval x
is the zero ideal. This is similar to algebraicIndependent_iff_ker_eq_bot
.
An element of R
is algebraic, when viewed as an element of the R
-algebra A
.
This is slightly more general than IsAlgebraic.algebraMap
in that it
allows noncommutative intermediate rings A
.
Transfer Algebra.IsAlgebraic
across an AlgEquiv
.
Alias of the reverse direction of IsAlgebraic.inv_iff
.
If x
is algebraic over R
, then x
is algebraic over S
when S
is an extension of R
,
and the map from R
to S
is injective.
Alias of IsAlgebraic.extendScalars
.
If x
is algebraic over R
, then x
is algebraic over S
when S
is an extension of R
,
and the map from R
to S
is injective.
A special case of IsAlgebraic.extendScalars
. This is extracted as a theorem
because in some cases IsAlgebraic.extendScalars
will just runs out of memory.
If x
is transcendental over S
, then x
is transcendental over R
when S
is an extension
of R
, and the map from R
to S
is injective.
Alias of Transcendental.restrictScalars
.
If x
is transcendental over S
, then x
is transcendental over R
when S
is an extension
of R
, and the map from R
to S
is injective.
A special case of Transcendental.restrictScalars
. This is extracted as a theorem
because in some cases Transcendental.restrictScalars
will just runs out of memory.
If A is an algebraic algebra over R, then A is algebraic over S when S is an extension of R,
and the map from R
to S
is injective.
Alias of Algebra.IsAlgebraic.extendScalars
.
If A is an algebraic algebra over R, then A is algebraic over S when S is an extension of R,
and the map from R
to S
is injective.
If x
is algebraic over K
, then x
is algebraic over L
when L
is an extension of K
Stacks Tag 09GF (part one)
If x
is transcendental over L
, then x
is transcendental over K
when
L
is an extension of K
If A is an algebraic algebra over K, then A is algebraic over L when L is an extension of K
Stacks Tag 09GF (part two)
Bijection between algebra equivalences and algebra homomorphisms
Equations
- Algebra.IsAlgebraic.algEquivEquivAlgHom K L = { toFun := fun (ϕ : L ≃ₐ[K] L) => ↑ϕ, invFun := fun (ϕ : L →ₐ[K] L) => AlgEquiv.ofBijective ϕ ⋯, left_inv := ⋯, right_inv := ⋯, map_mul' := ⋯ }
Instances For
A fraction (a : S) / (b : S)
can be reduced to (c : S) / (d : R)
,
if b
is algebraic over R
.
A fraction (a : S) / (b : S)
can be reduced to (c : S) / (d : R)
,
if b
is algebraic over R
.
In an algebraic extension L/K, an intermediate subalgebra is a field.