Documentation

Mathlib.RingTheory.Algebraic.Defs

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 #

Main results #

def IsAlgebraic (R : Type u) {A : Type v} [CommRing R] [Ring A] [Algebra R A] (x : A) :

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
Instances For
    def Transcendental (R : Type u) {A : Type v} [CommRing R] [Ring A] [Algebra R A] (x : A) :

    An element of an R-algebra is transcendental over R if it is not algebraic over R.

    Equations
    Instances For
      theorem transcendental_iff {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] {x : A} :
      Transcendental R x ∀ (p : Polynomial R), (Polynomial.aeval x) p = 0p = 0

      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.

      def Subalgebra.IsAlgebraic {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] (S : Subalgebra R A) :

      A subalgebra is algebraic if all its elements are algebraic.

      Equations
      Instances For
        class Algebra.IsAlgebraic (R : Type u) (A : Type v) [CommRing R] [Ring A] [Algebra R A] :

        An algebra is algebraic if all its elements are algebraic.

        Stacks Tag 09GC (Algebraic extensions)

        Instances
          class Algebra.Transcendental (R : Type u) (A : Type v) [CommRing R] [Ring A] [Algebra R A] :

          An algebra is transcendental if some element is transcendental.

          Instances
            theorem Algebra.isAlgebraic_def {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] :
            Algebra.IsAlgebraic R A ∀ (x : A), IsAlgebraic R x
            theorem Algebra.transcendental_def {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] :
            theorem Subalgebra.isAlgebraic_iff {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] (S : Subalgebra R A) :
            S.IsAlgebraic Algebra.IsAlgebraic R S

            A subalgebra is algebraic if and only if it is algebraic as an algebra.

            theorem Algebra.isAlgebraic_iff {R : Type u} {A : Type v} [CommRing R] [Ring A] [Algebra R A] :
            Algebra.IsAlgebraic R A .IsAlgebraic

            An algebra is algebraic if and only if it is algebraic as a subalgebra.