Documentation

Mathlib.FieldTheory.Adjoin

Adjoining Elements to Fields #

In this file we introduce the notion of adjoining elements to fields. This isn't quite the same as adjoining elements to rings. For example, Algebra.adjoin K {x} might not include x⁻¹.

Main results #

Notation #

def IntermediateField.adjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) :

adjoin F S extends a field F by adjoining a set S ⊆ E.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    @[simp]
    theorem IntermediateField.adjoin_le_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : Set E} {T : IntermediateField F E} :
    theorem IntermediateField.gc {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
    def IntermediateField.gi {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :

    Galois insertion between adjoin and coe.

    Equations
    • One or more equations did not get rendered due to their size.
    Instances For
      Equations
      Equations
      • IntermediateField.instInhabitedIntermediateField = { default := }
      theorem IntermediateField.coe_bot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
      theorem IntermediateField.mem_bot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {x : E} :
      @[simp]
      theorem IntermediateField.bot_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
      .toSubalgebra =
      @[simp]
      theorem IntermediateField.coe_top {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
      = Set.univ
      @[simp]
      theorem IntermediateField.mem_top {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {x : E} :
      @[simp]
      theorem IntermediateField.top_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
      .toSubalgebra =
      @[simp]
      theorem IntermediateField.coe_inf {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : IntermediateField F E) (T : IntermediateField F E) :
      ↑(S T) = S T
      @[simp]
      theorem IntermediateField.mem_inf {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : IntermediateField F E} {T : IntermediateField F E} {x : E} :
      x S T x S x T
      @[simp]
      theorem IntermediateField.inf_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : IntermediateField F E) (T : IntermediateField F E) :
      (S T).toSubalgebra = S.toSubalgebra T.toSubalgebra
      @[simp]
      theorem IntermediateField.coe_sInf {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set (IntermediateField F E)) :
      ↑(sInf S) = sInf ((fun x => x) '' S)
      @[simp]
      theorem IntermediateField.sInf_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set (IntermediateField F E)) :
      (sInf S).toSubalgebra = sInf (IntermediateField.toSubalgebra '' S)
      @[simp]
      theorem IntermediateField.sInf_toSubfield {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set (IntermediateField F E)) :
      IntermediateField.toSubfield (sInf S) = sInf (IntermediateField.toSubfield '' S)
      @[simp]
      theorem IntermediateField.coe_iInf {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Sort u_3} (S : ιIntermediateField F E) :
      ↑(iInf S) = ⋂ (i : ι), ↑(S i)
      @[simp]
      theorem IntermediateField.iInf_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Sort u_3} (S : ιIntermediateField F E) :
      (iInf S).toSubalgebra = ⨅ (i : ι), (S i).toSubalgebra
      @[simp]
      theorem IntermediateField.iInf_toSubfield {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Sort u_3} (S : ιIntermediateField F E) :
      @[simp]
      theorem IntermediateField.equivOfEq_apply {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : IntermediateField F E} {T : IntermediateField F E} (h : S = T) (x : { x // x S }) :
      ↑(IntermediateField.equivOfEq h) x = { val := x, property := (_ : x T) }
      def IntermediateField.equivOfEq {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : IntermediateField F E} {T : IntermediateField F E} (h : S = T) :
      { x // x S } ≃ₐ[F] { x // x T }

      Construct an algebra isomorphism from an equality of intermediate fields

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        @[simp]
        theorem IntermediateField.equivOfEq_rfl {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : IntermediateField F E) :
        IntermediateField.equivOfEq (_ : S = S) = AlgEquiv.refl
        @[simp]
        noncomputable def IntermediateField.botEquiv (F : Type u_1) [Field F] (E : Type u_2) [Field E] [Algebra F E] :
        { x // x } ≃ₐ[F] F

        The bottom intermediate_field is isomorphic to the field.

        Equations
        Instances For
          theorem IntermediateField.botEquiv_def {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (x : F) :
          ↑(IntermediateField.botEquiv F E) (↑(algebraMap F { x // x }) x) = x
          @[simp]
          theorem IntermediateField.botEquiv_symm {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (x : F) :
          ↑(AlgEquiv.symm (IntermediateField.botEquiv F E)) x = ↑(algebraMap F { x // x }) x
          noncomputable instance IntermediateField.algebraOverBot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
          Algebra { x // x } F
          Equations
          theorem IntermediateField.coe_algebraMap_over_bot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
          ↑(algebraMap { x // x } F) = ↑(IntermediateField.botEquiv F E)
          @[simp]
          theorem IntermediateField.topEquiv_symm_apply_coe {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
          ∀ (a : E), ↑(↑(AlgEquiv.symm IntermediateField.topEquiv) a) = ↑(↑(MulEquiv.symm Subalgebra.topEquiv) a)
          @[simp]
          theorem IntermediateField.topEquiv_apply {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
          ∀ (a : { x // x .toSubalgebra }), IntermediateField.topEquiv a = a
          def IntermediateField.topEquiv {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
          { x // x } ≃ₐ[F] E

          The top IntermediateField is isomorphic to the field.

          This is the intermediate field version of Subalgebra.topEquiv.

          Equations
          Instances For
            @[simp]
            theorem IntermediateField.restrictScalars_top {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : Type u_3} [Field K] [Algebra K E] [Algebra K F] [IsScalarTower K F E] :
            theorem IntermediateField.AlgHom.map_fieldRange {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : Type u_3} {L : Type u_4} [Field K] [Field L] [Algebra F K] [Algebra F L] (f : E →ₐ[F] K) (g : K →ₐ[F] L) :
            theorem IntermediateField.AlgHom.fieldRange_eq_top {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : Type u_3} [Field K] [Algebra F K] {f : E →ₐ[F] K} :
            @[simp]
            theorem IntermediateField.AlgEquiv.fieldRange_eq_top {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : Type u_3} [Field K] [Algebra F K] (f : E ≃ₐ[F] K) :
            theorem IntermediateField.adjoin.algebraMap_mem (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) (x : F) :
            instance IntermediateField.adjoin.fieldCoe (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) :
            Equations
            theorem IntermediateField.subset_adjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) :
            instance IntermediateField.adjoin.setCoe (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) :
            CoeTC S { x // x IntermediateField.adjoin F S }
            Equations
            theorem IntermediateField.adjoin.mono (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) (T : Set E) (h : S T) :
            theorem IntermediateField.subset_adjoin_of_subset_left {E : Type u_2} [Field E] (S : Set E) {F : Subfield E} {T : Set E} (HT : T F) :
            T ↑(IntermediateField.adjoin { x // x F } S)
            theorem IntermediateField.subset_adjoin_of_subset_right (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) {T : Set E} (H : T S) :
            @[simp]
            theorem IntermediateField.adjoin_empty (F : Type u_3) (E : Type u_4) [Field F] [Field E] [Algebra F E] :
            F⟮⟯ =
            @[simp]
            theorem IntermediateField.adjoin_univ (F : Type u_3) (E : Type u_4) [Field F] [Field E] [Algebra F E] :
            theorem IntermediateField.adjoin_le_subfield (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) {K : Subfield E} (HF : Set.range ↑(algebraMap F E) K) (HS : S K) :

            If K is a field with F ⊆ K and S ⊆ K then adjoin F S ≤ K.

            theorem IntermediateField.adjoin_subset_adjoin_iff (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {F' : Type u_3} [Field F'] [Algebra F' E] {S : Set E} {S' : Set E} :
            theorem IntermediateField.adjoin_map (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) {E' : Type u_3} [Field E'] [Algebra F E'] (f : E →ₐ[F] E') :
            theorem IntermediateField.algebra_adjoin_le_adjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) :
            theorem IntermediateField.adjoin_eq_algebra_adjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) (inv_mem : ∀ (x : E), x Algebra.adjoin F Sx⁻¹ Algebra.adjoin F S) :
            theorem IntermediateField.eq_adjoin_of_eq_algebra_adjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Set E) (K : IntermediateField F E) (h : K.toSubalgebra = Algebra.adjoin F S) :
            theorem IntermediateField.adjoin_induction (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {s : Set E} {p : EProp} {x : E} (h : x IntermediateField.adjoin F s) (Hs : (x : E) → x sp x) (Hmap : (x : F) → p (↑(algebraMap F E) x)) (Hadd : (x y : E) → p xp yp (x + y)) (Hneg : (x : E) → p xp (-x)) (Hinv : (x : E) → p xp x⁻¹) (Hmul : (x y : E) → p xp yp (x * y)) :
            p x

            If x₁ x₂ ... xₙ : E then F⟮x₁,x₂,...,xₙ⟯ is the IntermediateField F E generated by these elements.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For
              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                theorem IntermediateField.mem_adjoin_simple_self (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) :
                α Fα
                def IntermediateField.AdjoinSimple.gen (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) :
                { x // x Fα }

                generator of F⟮α⟯

                Equations
                Instances For
                  @[simp]
                  theorem IntermediateField.AdjoinSimple.algebraMap_gen (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) :
                  ↑(algebraMap { x // x Fα } E) (IntermediateField.AdjoinSimple.gen F α) = α
                  theorem IntermediateField.adjoin_simple_adjoin_simple (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) (β : E) :
                  IntermediateField.restrictScalars F { x // x Fα }β = Fα, β
                  theorem IntermediateField.adjoin_simple_comm (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) (β : E) :
                  IntermediateField.restrictScalars F { x // x Fα }β = IntermediateField.restrictScalars F { x // x Fβ }α
                  theorem IntermediateField.adjoin_algebraic_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : Set E} (hS : ∀ (x : E), x SIsAlgebraic F x) :
                  theorem IntermediateField.adjoin_simple_toSubalgebra_of_integral {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} (hα : IsIntegral F α) :
                  Fα.toSubalgebra = Algebra.adjoin F {α}
                  theorem IntermediateField.isSplittingField_iSup {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Type u_3} {t : ιIntermediateField F E} {p : ιPolynomial F} {s : Finset ι} (h0 : (Finset.prod s fun i => p i) 0) (h : ∀ (i : ι), i sPolynomial.IsSplittingField F { x // x t i } (p i)) :
                  Polynomial.IsSplittingField F { x // x ⨆ (i : ι) (_ : i s), t i } (Finset.prod s fun i => p i)

                  A compositum of splitting fields is a splitting field

                  theorem IntermediateField.adjoin_simple_le_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} {K : IntermediateField F E} :
                  Fα K α K

                  Adjoining a single element is compact in the lattice of intermediate fields.

                  Adjoining a finite subset is compact in the lattice of intermediate fields.

                  Adjoining a finite subset is compact in the lattice of intermediate fields.

                  The lattice of intermediate fields is compactly generated.

                  Equations
                  • One or more equations did not get rendered due to their size.
                  theorem IntermediateField.exists_finset_of_mem_iSup {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Type u_3} {f : ιIntermediateField F E} {x : E} (hx : x ⨆ (i : ι), f i) :
                  s, x ⨆ (i : ι) (_ : i s), f i
                  theorem IntermediateField.exists_finset_of_mem_supr' {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Type u_3} {f : ιIntermediateField F E} {x : E} (hx : x ⨆ (i : ι), f i) :
                  s, x ⨆ (i : (i : ι) × { x // x f i }) (_ : i s), Fi.snd
                  theorem IntermediateField.exists_finset_of_mem_supr'' {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {ι : Type u_3} {f : ιIntermediateField F E} (h : ∀ (i : ι), Algebra.IsAlgebraic F { x // x f i }) {x : E} (hx : x ⨆ (i : ι), f i) :
                  s, x ⨆ (i : (i : ι) × { x // x f i }) (_ : i s), IntermediateField.adjoin F (Polynomial.rootSet (minpoly F i.snd) E)
                  @[simp]
                  theorem IntermediateField.adjoin_eq_bot_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : Set E} :
                  theorem IntermediateField.adjoin_simple_eq_bot_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} :
                  Fα = α
                  @[simp]
                  theorem IntermediateField.adjoin_zero {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
                  F0 =
                  @[simp]
                  theorem IntermediateField.adjoin_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
                  F1 =
                  @[simp]
                  theorem IntermediateField.adjoin_int {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (n : ) :
                  Fn =
                  @[simp]
                  theorem IntermediateField.adjoin_nat {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (n : ) :
                  Fn =
                  @[simp]
                  theorem IntermediateField.rank_eq_one_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : IntermediateField F E} :
                  Module.rank F { x // x K } = 1 K =
                  @[simp]
                  theorem IntermediateField.finrank_eq_one_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {K : IntermediateField F E} :
                  FiniteDimensional.finrank F { x // x K } = 1 K =
                  @[simp]
                  theorem IntermediateField.rank_bot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
                  Module.rank F { x // x } = 1
                  @[simp]
                  theorem IntermediateField.finrank_bot {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] :
                  theorem IntermediateField.rank_adjoin_eq_one_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {S : Set E} :
                  theorem IntermediateField.rank_adjoin_simple_eq_one_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} :
                  Module.rank F { x // x Fα } = 1 α
                  theorem IntermediateField.finrank_adjoin_simple_eq_one_iff {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} :
                  FiniteDimensional.finrank F { x // x Fα } = 1 α
                  theorem IntermediateField.bot_eq_top_of_rank_adjoin_eq_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (h : ∀ (x : E), Module.rank F { x_1 // x_1 Fx } = 1) :

                  If F⟮x⟯ has dimension 1 over F for every x ∈ E then F = E.

                  theorem IntermediateField.bot_eq_top_of_finrank_adjoin_eq_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (h : ∀ (x : E), FiniteDimensional.finrank F { x_1 // x_1 Fx } = 1) :
                  theorem IntermediateField.subsingleton_of_rank_adjoin_eq_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (h : ∀ (x : E), Module.rank F { x_1 // x_1 Fx } = 1) :
                  theorem IntermediateField.subsingleton_of_finrank_adjoin_eq_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (h : ∀ (x : E), FiniteDimensional.finrank F { x_1 // x_1 Fx } = 1) :
                  theorem IntermediateField.bot_eq_top_of_finrank_adjoin_le_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] [FiniteDimensional F E] (h : ∀ (x : E), FiniteDimensional.finrank F { x_1 // x_1 Fx } 1) :

                  If F⟮x⟯ has dimension ≤1 over F for every x ∈ E then F = E.

                  theorem IntermediateField.subsingleton_of_finrank_adjoin_le_one {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] [FiniteDimensional F E] (h : ∀ (x : E), FiniteDimensional.finrank F { x_1 // x_1 Fx } 1) :
                  theorem IntermediateField.minpoly_gen (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] (α : E) :
                  noncomputable def IntermediateField.adjoinRootEquivAdjoin (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} (h : IsIntegral F α) :
                  AdjoinRoot (minpoly F α) ≃ₐ[F] { x // x Fα }

                  algebra isomorphism between AdjoinRoot and F⟮α⟯

                  Equations
                  • One or more equations did not get rendered due to their size.
                  Instances For
                    noncomputable def IntermediateField.powerBasisAux {K : Type u_3} [Field K] {L : Type u_4} [Field L] [Algebra K L] {x : L} (hx : IsIntegral K x) :
                    Basis (Fin (Polynomial.natDegree (minpoly K x))) K { x // x Kx }

                    The elements 1, x, ..., x ^ (d - 1) form a basis for K⟮x⟯, where d is the degree of the minimal polynomial of x.

                    Equations
                    Instances For
                      noncomputable def IntermediateField.adjoin.powerBasis {K : Type u_3} [Field K] {L : Type u_4} [Field L] [Algebra K L] {x : L} (hx : IsIntegral K x) :
                      PowerBasis K { x // x Kx }

                      The power basis 1, x, ..., x ^ (d - 1) for K⟮x⟯, where d is the degree of the minimal polynomial of x.

                      Equations
                      • One or more equations did not get rendered due to their size.
                      Instances For
                        theorem IntermediateField.adjoin.finiteDimensional {K : Type u_3} [Field K] {L : Type u_4} [Field L] [Algebra K L] {x : L} (hx : IsIntegral K x) :
                        FiniteDimensional K { x // x Kx }
                        theorem IntermediateField.adjoin.finrank {K : Type u_3} [Field K] {L : Type u_4} [Field L] [Algebra K L] {x : L} (hx : IsIntegral K x) :
                        theorem minpoly.degree_le {K : Type u_3} [Field K] {L : Type u_4} [Field L] [Algebra K L] (x : L) [FiniteDimensional K L] :
                        noncomputable def IntermediateField.algHomAdjoinIntegralEquiv (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} {K : Type u_3} [Field K] [Algebra F K] (h : IsIntegral F α) :
                        ({ x // x Fα } →ₐ[F] K) { x // x Polynomial.aroots (minpoly F α) K }

                        Algebra homomorphism F⟮α⟯ →ₐ[F] K are in bijection with the set of roots of minpoly α in K.

                        Equations
                        • One or more equations did not get rendered due to their size.
                        Instances For
                          noncomputable def IntermediateField.fintypeOfAlgHomAdjoinIntegral (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} {K : Type u_3} [Field K] [Algebra F K] (h : IsIntegral F α) :
                          Fintype ({ x // x Fα } →ₐ[F] K)

                          Fintype of algebra homomorphism F⟮α⟯ →ₐ[F] K

                          Equations
                          Instances For
                            theorem IntermediateField.card_algHom_adjoin_integral (F : Type u_1) [Field F] {E : Type u_2} [Field E] [Algebra F E] {α : E} {K : Type u_3} [Field K] [Algebra F K] (h : IsIntegral F α) (h_sep : Polynomial.Separable (minpoly F α)) (h_splits : Polynomial.Splits (algebraMap F K) (minpoly F α)) :
                            Fintype.card ({ x // x Fα } →ₐ[F] K) = Polynomial.natDegree (minpoly F α)
                            def IntermediateField.FG {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : IntermediateField F E) :

                            An intermediate field S is finitely generated if there exists t : Finset E such that IntermediateField.adjoin F t = S.

                            Equations
                            Instances For
                              theorem IntermediateField.fG_of_fG_toSubalgebra {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : IntermediateField F E) (h : Subalgebra.FG S.toSubalgebra) :
                              theorem IntermediateField.induction_on_adjoin_finset {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (S : Finset E) (P : IntermediateField F EProp) (base : P ) (ih : (K : IntermediateField F E) → (x : E) → x SP KP (IntermediateField.restrictScalars F { x // x K }x)) :
                              theorem IntermediateField.induction_on_adjoin_fg {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] (P : IntermediateField F EProp) (base : P ) (ih : (K : IntermediateField F E) → (x : E) → P KP (IntermediateField.restrictScalars F { x // x K }x)) (K : IntermediateField F E) (hK : IntermediateField.FG K) :
                              P K
                              theorem IntermediateField.induction_on_adjoin {F : Type u_1} [Field F] {E : Type u_2} [Field E] [Algebra F E] [FiniteDimensional F E] (P : IntermediateField F EProp) (base : P ) (ih : (K : IntermediateField F E) → (x : E) → P KP (IntermediateField.restrictScalars F { x // x K }x)) (K : IntermediateField F E) :
                              P K
                              def IntermediateField.Lifts (F : Type u_1) (E : Type u_2) (K : Type u_3) [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] :
                              Type (max u_2 u_3 u_2)

                              Lifts L → K of F → K

                              Equations
                              Instances For
                                instance IntermediateField.instPartialOrderLifts {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] :
                                Equations
                                noncomputable instance IntermediateField.instOrderBotLiftsToLEToPreorderInstPartialOrderLifts {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] :
                                Equations
                                • One or more equations did not get rendered due to their size.
                                noncomputable instance IntermediateField.instInhabitedLifts {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] :
                                Equations
                                • IntermediateField.instInhabitedLifts = { default := }
                                theorem IntermediateField.Lifts.eq_of_le {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {x : IntermediateField.Lifts F E K} {y : IntermediateField.Lifts F E K} (hxy : x y) (s : { x // x x.fst }) :
                                x.snd s = y.snd { val := s, property := (_ : s y.fst) }
                                theorem IntermediateField.Lifts.exists_max_two {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {c : Set (IntermediateField.Lifts F E K)} {x : IntermediateField.Lifts F E K} {y : IntermediateField.Lifts F E K} (hc : IsChain (fun x x_1 => x x_1) c) (hx : x insert c) (hy : y insert c) :
                                z, z insert c x z y z
                                theorem IntermediateField.Lifts.exists_max_three {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {c : Set (IntermediateField.Lifts F E K)} {x : IntermediateField.Lifts F E K} {y : IntermediateField.Lifts F E K} {z : IntermediateField.Lifts F E K} (hc : IsChain (fun x x_1 => x x_1) c) (hx : x insert c) (hy : y insert c) (hz : z insert c) :
                                w, w insert c x w y w z w
                                def IntermediateField.Lifts.upperBoundIntermediateField {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {c : Set (IntermediateField.Lifts F E K)} (hc : IsChain (fun x x_1 => x x_1) c) :

                                An upper bound on a chain of lifts

                                Equations
                                • One or more equations did not get rendered due to their size.
                                Instances For
                                  noncomputable def IntermediateField.Lifts.upperBoundAlgHom {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {c : Set (IntermediateField.Lifts F E K)} (hc : IsChain (fun x x_1 => x x_1) c) :

                                  The lift on the upper bound on a chain of lifts

                                  Equations
                                  • One or more equations did not get rendered due to their size.
                                  Instances For
                                    noncomputable def IntermediateField.Lifts.upperBound {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {c : Set (IntermediateField.Lifts F E K)} (hc : IsChain (fun x x_1 => x x_1) c) :

                                    An upper bound on a chain of lifts

                                    Equations
                                    Instances For
                                      theorem IntermediateField.Lifts.exists_upper_bound {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] (c : Set (IntermediateField.Lifts F E K)) (hc : IsChain (fun x x_1 => x x_1) c) :
                                      ub, ∀ (a : IntermediateField.Lifts F E K), a ca ub
                                      noncomputable def IntermediateField.Lifts.liftOfSplits {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] (x : IntermediateField.Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : Polynomial.Splits (algebraMap F K) (minpoly F s)) :

                                      Extend a lift x : Lifts F E K to an element s : E whose conjugates are all in K

                                      Equations
                                      • One or more equations did not get rendered due to their size.
                                      Instances For
                                        theorem IntermediateField.Lifts.le_lifts_of_splits {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] (x : IntermediateField.Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : Polynomial.Splits (algebraMap F K) (minpoly F s)) :
                                        theorem IntermediateField.Lifts.mem_lifts_of_splits {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] (x : IntermediateField.Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : Polynomial.Splits (algebraMap F K) (minpoly F s)) :
                                        theorem IntermediateField.Lifts.exists_lift_of_splits {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] (x : IntermediateField.Lifts F E K) {s : E} (h1 : IsIntegral F s) (h2 : Polynomial.Splits (algebraMap F K) (minpoly F s)) :
                                        y, x y s y.fst
                                        theorem IntermediateField.algHom_mk_adjoin_splits {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {S : Set E} (hK : ∀ (s : E), s SIsIntegral F s Polynomial.Splits (algebraMap F K) (minpoly F s)) :
                                        theorem IntermediateField.algHom_mk_adjoin_splits' {F : Type u_1} {E : Type u_2} {K : Type u_3} [Field F] [Field E] [Field K] [Algebra F E] [Algebra F K] {S : Set E} (hS : IntermediateField.adjoin F S = ) (hK : ∀ (x : E), x SIsIntegral F x Polynomial.Splits (algebraMap F K) (minpoly F x)) :
                                        theorem IntermediateField.le_sup_toSubalgebra {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] (E1 : IntermediateField K L) (E2 : IntermediateField K L) :
                                        E1.toSubalgebra E2.toSubalgebra (E1 E2).toSubalgebra
                                        theorem IntermediateField.sup_toSubalgebra {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] (E1 : IntermediateField K L) (E2 : IntermediateField K L) [h1 : FiniteDimensional K { x // x E1 }] [h2 : FiniteDimensional K { x // x E2 }] :
                                        (E1 E2).toSubalgebra = E1.toSubalgebra E2.toSubalgebra
                                        instance IntermediateField.finiteDimensional_iSup_of_finset {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] {ι : Type u_3} {f : ιIntermediateField K L} {s : Finset ι} [h : ∀ (i : ι), FiniteDimensional K { x // x f i }] :
                                        FiniteDimensional K { x // x ⨆ (i : ι) (_ : i s), f i }
                                        Equations
                                        theorem IntermediateField.finiteDimensional_iSup_of_finset' {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] {ι : Type u_3} {f : ιIntermediateField K L} {s : Finset ι} (h : ∀ (i : ι), i sFiniteDimensional K { x // x f i }) :
                                        FiniteDimensional K { x // x ⨆ (i : ι) (_ : i s), f i }
                                        theorem IntermediateField.isAlgebraic_iSup {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] {ι : Type u_3} {f : ιIntermediateField K L} (h : ∀ (i : ι), Algebra.IsAlgebraic K { x // x f i }) :
                                        Algebra.IsAlgebraic K { x // x ⨆ (i : ι), f i }

                                        A compositum of algebraic extensions is algebraic

                                        noncomputable def PowerBasis.equivAdjoinSimple {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] (pb : PowerBasis K L) :
                                        { x // x Kpb.gen } ≃ₐ[K] L

                                        pb.equivAdjoinSimple is the equivalence between K⟮pb.gen⟯ and L itself.

                                        Equations
                                        • One or more equations did not get rendered due to their size.
                                        Instances For
                                          @[simp]
                                          theorem PowerBasis.equivAdjoinSimple_aeval {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] (pb : PowerBasis K L) (f : Polynomial K) :
                                          @[simp]
                                          theorem PowerBasis.equivAdjoinSimple_gen {K : Type u_1} {L : Type u_2} [Field K] [Field L] [Algebra K L] (pb : PowerBasis K L) :