Documentation

Mathlib.FieldTheory.IsAlgClosed.Basic

Algebraically Closed Field #

In this file we define the typeclass for algebraically closed fields and algebraic closures, and prove some of their properties.

Main Definitions #

Tags #

algebraic closure, algebraically closed

class IsAlgClosed (k : Type u) [Field k] :

Typeclass for algebraically closed fields.

To show Polynomial.Splits p f for an arbitrary ring homomorphism f, see IsAlgClosed.splits_codomain and IsAlgClosed.splits_domain.

Instances
    theorem IsAlgClosed.splits_codomain {k : Type u_1} {K : Type u_2} [Field k] [IsAlgClosed k] [Field K] {f : K →+* k} (p : Polynomial K) :

    Every polynomial splits in the field extension f : K →+* k if k is algebraically closed.

    See also IsAlgClosed.splits_domain for the case where K is algebraically closed.

    theorem IsAlgClosed.splits_domain {k : Type u_1} {K : Type u_2} [Field k] [IsAlgClosed k] [Field K] {f : k →+* K} (p : Polynomial k) :

    Every polynomial splits in the field extension f : K →+* k if K is algebraically closed.

    See also IsAlgClosed.splits_codomain for the case where k is algebraically closed.

    theorem IsAlgClosed.exists_pow_nat_eq {k : Type u} [Field k] [IsAlgClosed k] (x : k) {n : } (hn : 0 < n) :
    z, z ^ n = x
    theorem IsAlgClosed.exists_eq_mul_self {k : Type u} [Field k] [IsAlgClosed k] (x : k) :
    z, x = z * z
    theorem IsAlgClosed.roots_eq_zero_iff {k : Type u} [Field k] [IsAlgClosed k] {p : Polynomial k} :
    Polynomial.roots p = 0 p = Polynomial.C (Polynomial.coeff p 0)
    theorem IsAlgClosed.exists_eval₂_eq_zero_of_injective {k : Type u} [Field k] {R : Type u_1} [Ring R] [IsAlgClosed k] (f : R →+* k) (hf : Function.Injective f) (p : Polynomial R) (hp : Polynomial.degree p 0) :
    x, Polynomial.eval₂ f x p = 0
    theorem IsAlgClosed.exists_eval₂_eq_zero {k : Type u} [Field k] {R : Type u_1} [Field R] [IsAlgClosed k] (f : R →+* k) (p : Polynomial R) (hp : Polynomial.degree p 0) :
    x, Polynomial.eval₂ f x p = 0
    theorem IsAlgClosed.exists_aeval_eq_zero_of_injective (k : Type u) [Field k] {R : Type u_1} [CommRing R] [IsAlgClosed k] [Algebra R k] (hinj : Function.Injective ↑(algebraMap R k)) (p : Polynomial R) (hp : Polynomial.degree p 0) :
    x, ↑(Polynomial.aeval x) p = 0
    theorem IsAlgClosed.exists_aeval_eq_zero (k : Type u) [Field k] {R : Type u_1} [Field R] [IsAlgClosed k] [Algebra R k] (p : Polynomial R) (hp : Polynomial.degree p 0) :
    x, ↑(Polynomial.aeval x) p = 0
    theorem IsAlgClosed.of_exists_root (k : Type u) [Field k] (H : ∀ (p : Polynomial k), Polynomial.Monic pIrreducible px, Polynomial.eval x p = 0) :
    theorem IsAlgClosed.of_ringEquiv (k : Type u) [Field k] (k' : Type u) [Field k'] (e : k ≃+* k') [IsAlgClosed k] :
    class IsAlgClosure (R : Type u) (K : Type v) [CommRing R] [Field K] [Algebra R K] [NoZeroSMulDivisors R K] :

    Typeclass for an extension being an algebraic closure.

    Instances
      structure IsAlgClosed.lift.SubfieldWithHom (K : Type u) (L : Type v) (M : Type w) [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] :
      Type (max v w)

      This structure is used to prove the existence of a homomorphism from any algebraic extension into an algebraic closure

      Instances For
        Equations
        • IsAlgClosed.lift.SubfieldWithHom.instLESubfieldWithHom = { le := fun E₁ E₂ => h, ∀ (x : { x // x E₁.carrier }), E₂.emb (↑(Subalgebra.inclusion h) x) = E₁.emb x }
        Equations
        theorem IsAlgClosed.lift.SubfieldWithHom.le_def {K : Type u} {L : Type v} {M : Type w} [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] {E₁ : IsAlgClosed.lift.SubfieldWithHom K L M} {E₂ : IsAlgClosed.lift.SubfieldWithHom K L M} :
        E₁ E₂ h, ∀ (x : { x // x E₁.carrier }), E₂.emb (↑(Subalgebra.inclusion h) x) = E₁.emb x
        theorem IsAlgClosed.lift.SubfieldWithHom.compat {K : Type u} {L : Type v} {M : Type w} [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] {E₁ : IsAlgClosed.lift.SubfieldWithHom K L M} {E₂ : IsAlgClosed.lift.SubfieldWithHom K L M} (h : E₁ E₂) (x : { x // x E₁.carrier }) :
        E₂.emb (↑(Subalgebra.inclusion (_ : E₁.carrier E₂.carrier)) x) = E₁.emb x
        Equations
        • One or more equations did not get rendered due to their size.
        theorem IsAlgClosed.lift.SubfieldWithHom.maximal_subfieldWithHom_chain_bounded {K : Type u} {L : Type v} {M : Type w} [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] (c : Set (IsAlgClosed.lift.SubfieldWithHom K L M)) (hc : IsChain (fun x x_1 => x x_1) c) :
        ub, ∀ (N : IsAlgClosed.lift.SubfieldWithHom K L M), N cN ub
        theorem IsAlgClosed.lift.SubfieldWithHom.exists_maximal_subfieldWithHom (K : Type u) (L : Type v) (M : Type w) [Field K] [Field L] [Algebra K L] [Field M] [Algebra K M] :
        E, ∀ (N : IsAlgClosed.lift.SubfieldWithHom K L M), E NN E

        The maximal SubfieldWithHom. We later prove that this is equal to .

        Equations
        Instances For
          noncomputable def Algebra.adjoin.liftSingleton (R : Type u_1) [Field R] {S : Type u_2} {T : Type u_3} [CommRing S] [CommRing T] [Algebra R S] [Algebra R T] (x : S) (y : T) (h : ↑(Polynomial.aeval y) (minpoly R x) = 0) :
          { x // x Algebra.adjoin R {x} } →ₐ[R] T

          Produce an algebra homomorphism Adjoin R {x} →ₐ[R] T sending x to a root of x's minimal polynomial in T.

          Equations
          Instances For
            @[irreducible]
            noncomputable def IsAlgClosed.lift {M : Type u_1} [Field M] [IsAlgClosed M] {R : Type u_2} [CommRing R] {S : Type u_3} [CommRing S] [IsDomain S] [Algebra R S] [Algebra R M] [NoZeroSMulDivisors R S] [NoZeroSMulDivisors R M] (hS : Algebra.IsAlgebraic R S) :

            A (random) homomorphism from an algebraic extension of R into an algebraically closed extension of R.

            Equations
            Instances For
              theorem IsAlgClosed.lift_def {M : Type u_1} [Field M] [IsAlgClosed M] {R : Type u_2} [CommRing R] {S : Type u_3} [CommRing S] [IsDomain S] [Algebra R S] [Algebra R M] [NoZeroSMulDivisors R S] [NoZeroSMulDivisors R M] (hS : Algebra.IsAlgebraic R S) :
              IsAlgClosed.lift hS = let_fun this := (_ : IsScalarTower R (FractionRing R) M); let_fun this_1 := (_ : IsScalarTower R (FractionRing R) (FractionRing S)); let_fun this_2 := (_ : Algebra.IsAlgebraic (FractionRing R) (FractionRing S)); let f := IsAlgClosed.lift_aux (FractionRing R) (FractionRing S) M this_2; AlgHom.comp (AlgHom.restrictScalars R f) (AlgHom.restrictScalars R (Algebra.ofId S (FractionRing S)))

              Algebraically closed fields are infinite since Xⁿ⁺¹ - 1 is separable when #K = n

              Equations
              noncomputable def IsAlgClosure.equiv (R : Type u) [CommRing R] (L : Type v) (M : Type w) [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra R L] [NoZeroSMulDivisors R L] [IsAlgClosure R L] :

              A (random) isomorphism between two algebraic closures of R.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                theorem IsAlgClosure.ofAlgebraic (K : Type u_1) (J : Type u_2) (L : Type v) [Field K] [Field J] [Field L] [Algebra K J] [Algebra J L] [IsAlgClosure J L] [Algebra K L] [IsScalarTower K J L] (hKJ : Algebra.IsAlgebraic K J) :

                If J is an algebraic extension of K and L is an algebraic closure of J, then it is also an algebraic closure of K.

                noncomputable def IsAlgClosure.equivOfAlgebraic' (R : Type u) (S : Type u_3) (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] [Algebra R S] [Algebra R L] [IsScalarTower R S L] [Nontrivial S] [NoZeroSMulDivisors R S] (hRL : Algebra.IsAlgebraic R L) :

                A (random) isomorphism between an algebraic closure of R and an algebraic closure of an algebraic extension of R

                Equations
                Instances For
                  noncomputable def IsAlgClosure.equivOfAlgebraic (K : Type u_1) (J : Type u_2) (L : Type v) (M : Type w) [Field K] [Field J] [Field L] [Field M] [Algebra K M] [IsAlgClosure K M] [Algebra K J] [Algebra J L] [IsAlgClosure J L] [Algebra K L] [IsScalarTower K J L] (hKJ : Algebra.IsAlgebraic K J) :

                  A (random) isomorphism between an algebraic closure of K and an algebraic closure of an algebraic extension of K

                  Equations
                  Instances For
                    noncomputable def IsAlgClosure.equivOfEquivAux {R : Type u} {S : Type u_3} (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] (hSR : S ≃+* R) :

                    Used in the definition of equivOfEquiv

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      noncomputable def IsAlgClosure.equivOfEquiv {R : Type u} {S : Type u_3} (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] (hSR : S ≃+* R) :
                      L ≃+* M

                      Algebraic closure of isomorphic fields are isomorphic

                      Equations
                      Instances For
                        @[simp]
                        theorem IsAlgClosure.equivOfEquiv_comp_algebraMap {R : Type u} {S : Type u_3} (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] (hSR : S ≃+* R) :
                        @[simp]
                        theorem IsAlgClosure.equivOfEquiv_algebraMap {R : Type u} {S : Type u_3} (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] (hSR : S ≃+* R) (s : S) :
                        ↑(IsAlgClosure.equivOfEquiv L M hSR) (↑(algebraMap S L) s) = ↑(algebraMap R M) (hSR s)
                        @[simp]
                        theorem IsAlgClosure.equivOfEquiv_symm_algebraMap {R : Type u} {S : Type u_3} (L : Type v) (M : Type w) [CommRing R] [CommRing S] [Field L] [Field M] [Algebra R M] [NoZeroSMulDivisors R M] [IsAlgClosure R M] [Algebra S L] [NoZeroSMulDivisors S L] [IsAlgClosure S L] (hSR : S ≃+* R) (r : R) :
                        ↑(RingEquiv.symm (IsAlgClosure.equivOfEquiv L M hSR)) (↑(algebraMap R M) r) = ↑(algebraMap S L) (↑(RingEquiv.symm hSR) r)
                        theorem Algebra.IsAlgebraic.range_eval_eq_rootSet_minpoly {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [IsAlgClosed A] [Algebra F K] (hK : Algebra.IsAlgebraic F K) [Algebra F A] (x : K) :
                        (Set.range fun ψ => ψ x) = Polynomial.rootSet (minpoly F x) A

                        Let A be an algebraically closed field and let x ∈ K, with K/F an algebraic extension of fields. Then the images of x by the F-algebra morphisms from K to A are exactly the roots in A of the minimal polynomial of x over F.