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

Main results #

structure 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 For
    theorem IsAlgClosed.splits_codomain {k : Type u_1} {K : Type u_2} [Field k] [IsAlgClosed k] [Field K] {f : RingHom 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 : RingHom 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_root {k : Type u} [Field k] [IsAlgClosed k] (p : Polynomial k) (hp : Ne p.degree 0) :
    Exists fun (x : k) => p.IsRoot x

    If k is algebraically closed, then every nonconstant polynomial has a root.

    Stacks Tag 09GR ((4) ⟹ (3))

    theorem IsAlgClosed.exists_pow_nat_eq {k : Type u} [Field k] [IsAlgClosed k] (x : k) {n : Nat} (hn : LT.lt 0 n) :
    Exists fun (z : k) => Eq (HPow.hPow z n) x
    theorem IsAlgClosed.exists_eq_mul_self {k : Type u} [Field k] [IsAlgClosed k] (x : k) :
    Exists fun (z : k) => Eq x (HMul.hMul z z)
    theorem IsAlgClosed.exists_eval₂_eq_zero_of_injective {k : Type u} [Field k] {R : Type u_1} [Ring R] [IsAlgClosed k] (f : RingHom R k) (hf : Function.Injective (DFunLike.coe f)) (p : Polynomial R) (hp : Ne p.degree 0) :
    Exists fun (x : k) => Eq (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 : RingHom R k) (p : Polynomial R) (hp : Ne p.degree 0) :
    Exists fun (x : k) => Eq (Polynomial.eval₂ f 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 : Ne p.degree 0) :
    Exists fun (x : k) => Eq (DFunLike.coe (Polynomial.aeval x) p) 0
    theorem IsAlgClosed.of_exists_root (k : Type u) [Field k] (H : ∀ (p : Polynomial k), p.MonicIrreducible pExists fun (x : k) => Eq (Polynomial.eval x p) 0) :

    If every nonconstant polynomial over k has a root, then k is algebraically closed.

    Stacks Tag 09GR ((3) ⟹ (4))

    theorem IsAlgClosed.of_ringEquiv (k : Type u) [Field k] (k' : Type u) [Field k'] (e : RingEquiv k k') [IsAlgClosed k] :

    If k is algebraically closed, then every irreducible polynomial over k is linear.

    Stacks Tag 09GR ((4) ⟹ (2))

    If k is algebraically closed, K / k is a field extension, L / k is an intermediate field which is algebraic, then L is equal to k. A corollary of IsAlgClosed.algebraMap_surjective_of_isAlgebraic.

    Stacks Tag 09GQ (The result is the definition of algebraically closedness in Stacks Project. This statement is 09GR (4) ⟹ (1).)

    theorem Polynomial.isCoprime_iff_aeval_ne_zero_of_isAlgClosed (k : Type u) [Field k] (K : Type v) [Field K] [IsAlgClosed K] [Algebra k K] (p q : Polynomial k) :
    Iff (IsCoprime p q) (∀ (a : K), Or (Ne (DFunLike.coe (aeval a) p) 0) (Ne (DFunLike.coe (aeval a) q) 0))
    structure 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 For
      theorem isAlgClosure_iff (k : Type u) [Field k] (K : Type v) [Field K] [Algebra k K] :
      theorem IsAlgClosure.normal (R : Type u_1) (K : Type u_2) [Field R] [Field K] [Algebra R K] [IsAlgClosure R K] :
      Normal R K
      theorem IsAlgClosure.separable (R : Type u_1) (K : Type u_2) [Field R] [Field K] [Algebra R K] [IsAlgClosure R K] [CharZero R] :
      theorem IsAlgClosure.of_splits {R : Type u_1} {K : Type u_2} [CommRing R] [IsDomain R] [Field K] [Algebra R K] [Algebra.IsIntegral R K] [NoZeroSMulDivisors R K] (h : ∀ (p : Polynomial R), p.MonicIrreducible pPolynomial.Splits (algebraMap R K) p) :
      theorem IsAlgClosed.surjective_restrictDomain_of_isAlgebraic {K : Type u} [Field K] {L : Type v} {M : Type w} [Field L] [Algebra K L] [Field M] [Algebra K M] [IsAlgClosed M] {E : Type u_1} [Field E] [Algebra K E] [Algebra L E] [IsScalarTower K L E] [Algebra.IsAlgebraic L E] :

      If E/L/K is a tower of field extensions with E/L algebraic, and if M is an algebraically closed extension of K, then any embedding of L/K into M/K extends to an embedding of E/K. Known as the extension lemma in https://math.stackexchange.com/a/687914.

      @[deprecated IsAlgClosed.surjective_restrictDomain_of_isAlgebraic (since := "2024-11-15")]
      theorem IsAlgClosed.surjective_comp_algebraMap_of_isAlgebraic {K : Type u} [Field K] {L : Type v} {M : Type w} [Field L] [Algebra K L] [Field M] [Algebra K M] [IsAlgClosed M] {E : Type u_1} [Field E] [Algebra K E] [Algebra L E] [IsScalarTower K L E] [Algebra.IsAlgebraic L E] :

      Alias of IsAlgClosed.surjective_restrictDomain_of_isAlgebraic.


      If E/L/K is a tower of field extensions with E/L algebraic, and if M is an algebraically closed extension of K, then any embedding of L/K into M/K extends to an embedding of E/K. Known as the extension lemma in https://math.stackexchange.com/a/687914.

      @[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] [Algebra.IsAlgebraic R S] :
      AlgHom R S M

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

      Stacks Tag 09GU

      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] [Algebra.IsAlgebraic R S] :
        Eq lift (let_fun this := ; let_fun this_1 := ; let f := IsAlgClosed.lift_aux✝ (FractionRing R) (FractionRing S) M; (AlgHom.restrictScalars R f).comp (AlgHom.restrictScalars R (Algebra.ofId S (FractionRing S))))
        theorem IsAlgClosed.perfectRing (k : Type u) [Field k] (p : Nat) [Fact (Nat.Prime p)] [CharP k p] [IsAlgClosed k] :

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

        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] :
        AlgEquiv R L M

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

        Stacks Tag 09GV

        Equations
        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] [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] [Algebra.IsAlgebraic R L] :
          AlgEquiv R L M

          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] [Algebra.IsAlgebraic K J] :
            AlgEquiv K L M

            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 : RingEquiv S R) :
              Subtype fun (e : RingEquiv L M) => Eq (e.toRingHom.comp (algebraMap S L)) ((algebraMap R M).comp hSR.toRingHom)

              Used in the definition of equivOfEquiv

              Equations
              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 : RingEquiv S R) :

                Algebraic closure of isomorphic fields are isomorphic

                Equations
                Instances For
                  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 : RingEquiv S R) (s : S) :
                  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 : RingEquiv S R) (r : 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] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] [IsAlgClosed A] (x : K) :
                  Eq (Set.range fun (ψ : AlgHom F K A) => DFunLike.coe ψ x) ((minpoly F x).rootSet 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.

                  def IntermediateField.algHomEquivAlgHomOfSplits {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : IntermediateField F A) (hL : ∀ (x : K), Polynomial.Splits (algebraMap F (Subtype fun (x : A) => Membership.mem L x)) (minpoly F x)) :
                  Equiv (AlgHom F K (Subtype fun (x : A) => Membership.mem L x)) (AlgHom F K A)

                  All F-embeddings of a field K into another field A factor through any intermediate field of A/F in which the minimal polynomial of elements of K splits.

                  Equations
                  Instances For
                    theorem IntermediateField.algHomEquivAlgHomOfSplits_symm_apply {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : IntermediateField F A) (hL : ∀ (x : K), Polynomial.Splits (algebraMap F (Subtype fun (x : A) => Membership.mem L x)) (minpoly F x)) (f : AlgHom F K A) :
                    theorem IntermediateField.algHomEquivAlgHomOfSplits_apply {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : IntermediateField F A) (hL : ∀ (x : K), Polynomial.Splits (algebraMap F (Subtype fun (x : A) => Membership.mem L x)) (minpoly F x)) (φ₂ : AlgHom F K (Subtype fun (x : A) => Membership.mem L x)) :
                    Eq (DFunLike.coe (algHomEquivAlgHomOfSplits A L hL) φ₂) (L.val.comp φ₂)
                    theorem IntermediateField.algHomEquivAlgHomOfSplits_apply_apply {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : IntermediateField F A) (hL : ∀ (x : K), Polynomial.Splits (algebraMap F (Subtype fun (x : A) => Membership.mem L x)) (minpoly F x)) (f : AlgHom F K (Subtype fun (x : A) => Membership.mem L x)) (x : K) :
                    noncomputable def Algebra.IsAlgebraic.algHomEquivAlgHomOfSplits {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : Type u_4) [Field L] [Algebra F L] [Algebra L A] [IsScalarTower F L A] (hL : ∀ (x : K), Polynomial.Splits (algebraMap F L) (minpoly F x)) :
                    Equiv (AlgHom F K L) (AlgHom F K A)

                    All F-embeddings of a field K into another field A factor through any subextension of A/F in which the minimal polynomial of elements of K splits.

                    Equations
                    • One or more equations did not get rendered due to their size.
                    Instances For
                      theorem Algebra.IsAlgebraic.algHomEquivAlgHomOfSplits_apply_apply {F : Type u_1} {K : Type u_2} (A : Type u_3) [Field F] [Field K] [Field A] [Algebra F K] [Algebra F A] [Algebra.IsAlgebraic F K] (L : Type u_4) [Field L] [Algebra F L] [Algebra L A] [IsScalarTower F L A] (hL : ∀ (x : K), Polynomial.Splits (algebraMap F L) (minpoly F x)) (f : AlgHom F K L) (x : K) :
                      theorem Polynomial.isRoot_of_isRoot_iff_dvd_derivative_mul {K : Type u_1} [Field K] [IsAlgClosed K] [CharZero K] {f g : Polynomial K} (hf0 : Ne f 0) :
                      Iff (∀ (x : K), f.IsRoot xg.IsRoot x) (Dvd.dvd f (HMul.hMul (DFunLike.coe derivative f) g))

                      Over an algebraically closed field of characteristic zero a necessary and sufficient condition for the set of roots of a nonzero polynomial f to be a subset of the set of roots of g is that f divides f.derivative * g. Over an integral domain, this is a sufficient but not necessary condition. See isRoot_of_isRoot_of_dvd_derivative_mul