Documentation

Mathlib.FieldTheory.Subfield

Subfields #

Let K be a division ring, for example a field. This file defines the "bundled" subfield type Subfield K, a type whose terms correspond to subfields of K. Note we do not require the "subfields" to be commutative, so they are really sub-division rings / skew fields. This is the preferred way to talk about subfields in mathlib. Unbundled subfields (s : Set K and IsSubfield s) are not in this file, and they will ultimately be deprecated.

We prove that subfields are a complete lattice, and that you can map (pushforward) and comap (pull back) them along ring homomorphisms.

We define the closure construction from Set K to Subfield K, sending a subset of K to the subfield it generates, and prove that it is a Galois insertion.

Main definitions #

Notation used here:

(K : Type u) [DivisionRing K] (L : Type u) [DivisionRing L] (f g : K →+* L) (A : Subfield K) (B : Subfield L) (s : Set K)

Implementation notes #

A subfield is implemented as a subring which is closed under ⁻¹.

Lattice inclusion (e.g. and ) is used rather than set notation ( and ), although is defined as membership of a subfield's underlying set.

Tags #

subfield, subfields

class SubfieldClass (S : Type u_1) (K : Type u_2) [DivisionRing K] [SetLike S K] extends SubringClass , InvMemClass :

SubfieldClass S K states S is a type of subsets s ⊆ K closed under field operations.

    Instances
      instance SubfieldClass.toSubgroupClass {K : Type u} [DivisionRing K] (S : Type u_1) [SetLike S K] [h : SubfieldClass S K] :

      A subfield contains 1, products and inverses.

      Be assured that we're not actually proving that subfields are subgroups: SubgroupClass is really an abbreviation of SubgroupWithOrWithoutZeroClass.

      Equations
      • =
      theorem SubfieldClass.ratCast_mem {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (q : ) :
      q s
      instance SubfieldClass.instRatCastSubtypeMemInstMembership {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) :
      RatCast s
      Equations
      @[simp]
      theorem SubfieldClass.coe_ratCast {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (x : ) :
      x = x
      @[deprecated SubfieldClass.coe_ratCast]
      theorem SubfieldClass.coe_rat_cast {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (x : ) :
      x = x

      Alias of SubfieldClass.coe_ratCast.

      @[deprecated SubfieldClass.ratCast_mem]
      theorem SubfieldClass.coe_rat_mem {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (q : ) :
      q s

      Alias of SubfieldClass.ratCast_mem.

      theorem SubfieldClass.rat_smul_mem {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (a : ) (x : s) :
      a x s
      theorem SubfieldClass.ofScientific_mem {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) {b : Bool} {n : } {m : } :
      instance SubfieldClass.instSMulRatSubtypeMemInstMembership {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) :
      SMul s
      Equations
      @[simp]
      theorem SubfieldClass.coe_rat_smul {K : Type u} [DivisionRing K] {S : Type u_1} [SetLike S K] [h : SubfieldClass S K] (s : S) (a : ) (x : s) :
      (a x) = a x
      instance SubfieldClass.toDivisionRing {K : Type u} [DivisionRing K] (S : Type u_1) [SetLike S K] [h : SubfieldClass S K] (s : S) :

      A subfield inherits a division ring structure

      Equations
      instance SubfieldClass.toField (S : Type u_1) {K : Type u_2} [Field K] [SetLike S K] [SubfieldClass S K] (s : S) :
      Field s

      A subfield of a field inherits a field structure

      Equations
      structure Subfield (K : Type u) [DivisionRing K] extends Subring :

      Subfield R is the type of subfields of R. A subfield of R is a subset s that is a multiplicative submonoid and an additive subgroup. Note in particular that it shares the same 0 and 1 as R.

      • carrier : Set K
      • mul_mem' : ∀ {a b : K}, a self.carrierb self.carriera * b self.carrier
      • one_mem' : 1 self.carrier
      • add_mem' : ∀ {a b : K}, a self.carrierb self.carriera + b self.carrier
      • zero_mem' : 0 self.carrier
      • neg_mem' : ∀ {x : K}, x self.carrier-x self.carrier
      • inv_mem' : xself.carrier, x⁻¹ self.carrier

        A subfield is closed under multiplicative inverses.

      Instances For

        The underlying AddSubgroup of a subfield.

        Equations
        Instances For
          Equations
          • Subfield.instSetLikeSubfield = { coe := fun (s : Subfield K) => s.carrier, coe_injective' := }
          theorem Subfield.mem_carrier {K : Type u} [DivisionRing K] {s : Subfield K} {x : K} :
          x s.carrier x s
          @[simp]
          theorem Subfield.mem_mk {K : Type u} [DivisionRing K] {S : Subring K} {x : K} (h : xS.carrier, x⁻¹ S.carrier) :
          x { toSubring := S, inv_mem' := h } x S
          @[simp]
          theorem Subfield.coe_set_mk {K : Type u} [DivisionRing K] (S : Subring K) (h : xS.carrier, x⁻¹ S.carrier) :
          { toSubring := S, inv_mem' := h } = S
          @[simp]
          theorem Subfield.mk_le_mk {K : Type u} [DivisionRing K] {S : Subring K} {S' : Subring K} (h : xS.carrier, x⁻¹ S.carrier) (h' : xS'.carrier, x⁻¹ S'.carrier) :
          { toSubring := S, inv_mem' := h } { toSubring := S', inv_mem' := h' } S S'
          theorem Subfield.ext {K : Type u} [DivisionRing K] {S : Subfield K} {T : Subfield K} (h : ∀ (x : K), x S x T) :
          S = T

          Two subfields are equal if they have the same elements.

          def Subfield.copy {K : Type u} [DivisionRing K] (S : Subfield K) (s : Set K) (hs : s = S) :

          Copy of a subfield with a new carrier equal to the old one. Useful to fix definitional equalities.

          Equations
          • One or more equations did not get rendered due to their size.
          Instances For
            @[simp]
            theorem Subfield.coe_copy {K : Type u} [DivisionRing K] (S : Subfield K) (s : Set K) (hs : s = S) :
            (Subfield.copy S s hs) = s
            theorem Subfield.copy_eq {K : Type u} [DivisionRing K] (S : Subfield K) (s : Set K) (hs : s = S) :
            Subfield.copy S s hs = S
            @[simp]
            theorem Subfield.coe_toSubring {K : Type u} [DivisionRing K] (s : Subfield K) :
            s.toSubring = s
            @[simp]
            theorem Subfield.mem_toSubring {K : Type u} [DivisionRing K] (s : Subfield K) (x : K) :
            x s.toSubring x s
            def Subring.toSubfield {K : Type u} [DivisionRing K] (s : Subring K) (hinv : xs, x⁻¹ s) :

            A Subring containing inverses is a Subfield.

            Equations
            Instances For
              theorem Subfield.one_mem {K : Type u} [DivisionRing K] (s : Subfield K) :
              1 s

              A subfield contains the field's 1.

              theorem Subfield.zero_mem {K : Type u} [DivisionRing K] (s : Subfield K) :
              0 s

              A subfield contains the field's 0.

              theorem Subfield.mul_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} {y : K} :
              x sy sx * y s

              A subfield is closed under multiplication.

              theorem Subfield.add_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} {y : K} :
              x sy sx + y s

              A subfield is closed under addition.

              theorem Subfield.neg_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} :
              x s-x s

              A subfield is closed under negation.

              theorem Subfield.sub_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} {y : K} :
              x sy sx - y s

              A subfield is closed under subtraction.

              theorem Subfield.inv_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} :
              x sx⁻¹ s

              A subfield is closed under inverses.

              theorem Subfield.div_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} {y : K} :
              x sy sx / y s

              A subfield is closed under division.

              theorem Subfield.list_prod_mem {K : Type u} [DivisionRing K] (s : Subfield K) {l : List K} :
              (xl, x s)List.prod l s

              Product of a list of elements in a subfield is in the subfield.

              theorem Subfield.list_sum_mem {K : Type u} [DivisionRing K] (s : Subfield K) {l : List K} :
              (xl, x s)List.sum l s

              Sum of a list of elements in a subfield is in the subfield.

              theorem Subfield.multiset_sum_mem {K : Type u} [DivisionRing K] (s : Subfield K) (m : Multiset K) :
              (am, a s)Multiset.sum m s

              Sum of a multiset of elements in a Subfield is in the Subfield.

              theorem Subfield.sum_mem {K : Type u} [DivisionRing K] (s : Subfield K) {ι : Type u_1} {t : Finset ι} {f : ιK} (h : ct, f c s) :
              (Finset.sum t fun (i : ι) => f i) s

              Sum of elements in a Subfield indexed by a Finset is in the Subfield.

              theorem Subfield.pow_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} (hx : x s) (n : ) :
              x ^ n s
              theorem Subfield.zsmul_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} (hx : x s) (n : ) :
              n x s
              theorem Subfield.intCast_mem {K : Type u} [DivisionRing K] (s : Subfield K) (n : ) :
              n s
              @[deprecated intCast_mem]
              theorem Subfield.coe_int_mem {R : Type u} {S : Type v} [Ring R] [SetLike S R] [hSR : SubringClass S R] (s : S) (n : ) :
              n s

              Alias of intCast_mem.

              theorem Subfield.zpow_mem {K : Type u} [DivisionRing K] (s : Subfield K) {x : K} (hx : x s) (n : ) :
              x ^ n s
              Equations
              Equations
              Equations
              instance Subfield.toField {K : Type u_1} [Field K] (s : Subfield K) :
              Field s

              A subfield inherits a field structure

              Equations
              @[simp]
              theorem Subfield.coe_add {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) (y : s) :
              (x + y) = x + y
              @[simp]
              theorem Subfield.coe_sub {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) (y : s) :
              (x - y) = x - y
              @[simp]
              theorem Subfield.coe_neg {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) :
              (-x) = -x
              @[simp]
              theorem Subfield.coe_mul {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) (y : s) :
              (x * y) = x * y
              @[simp]
              theorem Subfield.coe_div {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) (y : s) :
              (x / y) = x / y
              @[simp]
              theorem Subfield.coe_inv {K : Type u} [DivisionRing K] (s : Subfield K) (x : s) :
              x⁻¹ = (x)⁻¹
              @[simp]
              theorem Subfield.coe_zero {K : Type u} [DivisionRing K] (s : Subfield K) :
              0 = 0
              @[simp]
              theorem Subfield.coe_one {K : Type u} [DivisionRing K] (s : Subfield K) :
              1 = 1
              def Subfield.subtype {K : Type u} [DivisionRing K] (s : Subfield K) :
              s →+* K

              The embedding from a subfield of the field K to K.

              Equations
              • One or more equations did not get rendered due to their size.
              Instances For
                @[simp]
                theorem Subfield.coe_subtype {K : Type u} [DivisionRing K] (s : Subfield K) :
                (Subfield.subtype s) = Subtype.val

                Partial order #

                theorem Subfield.mem_toSubmonoid {K : Type u} [DivisionRing K] {s : Subfield K} {x : K} :
                x s.toSubmonoid x s
                @[simp]
                theorem Subfield.coe_toSubmonoid {K : Type u} [DivisionRing K] (s : Subfield K) :
                s.toSubmonoid = s
                @[simp]
                @[simp]

                top #

                The subfield of K containing all elements of K.

                Equations
                • Subfield.instTopSubfield = { top := let __src := ; { toSubring := __src, inv_mem' := } }
                Equations
                • Subfield.instInhabitedSubfield = { default := }
                @[simp]
                theorem Subfield.mem_top {K : Type u} [DivisionRing K] (x : K) :
                @[simp]
                theorem Subfield.coe_top {K : Type u} [DivisionRing K] :
                = Set.univ

                The ring equiv between the top element of Subfield K and K.

                Equations
                • Subfield.topEquiv = Subsemiring.topEquiv
                Instances For

                  comap #

                  def Subfield.comap {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) (s : Subfield L) :

                  The preimage of a subfield along a ring homomorphism is a subfield.

                  Equations
                  Instances For
                    @[simp]
                    theorem Subfield.coe_comap {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) (s : Subfield L) :
                    (Subfield.comap f s) = f ⁻¹' s
                    @[simp]
                    theorem Subfield.mem_comap {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {s : Subfield L} {f : K →+* L} {x : K} :
                    theorem Subfield.comap_comap {K : Type u} {L : Type v} {M : Type w} [DivisionRing K] [DivisionRing L] [DivisionRing M] (s : Subfield M) (g : L →+* M) (f : K →+* L) :

                    map #

                    def Subfield.map {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) (s : Subfield K) :

                    The image of a subfield along a ring homomorphism is a subfield.

                    Equations
                    Instances For
                      @[simp]
                      theorem Subfield.coe_map {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (s : Subfield K) (f : K →+* L) :
                      (Subfield.map f s) = f '' s
                      @[simp]
                      theorem Subfield.mem_map {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {f : K →+* L} {s : Subfield K} {y : L} :
                      y Subfield.map f s ∃ x ∈ s, f x = y
                      theorem Subfield.map_map {K : Type u} {L : Type v} {M : Type w} [DivisionRing K] [DivisionRing L] [DivisionRing M] (s : Subfield K) (g : L →+* M) (f : K →+* L) :
                      theorem Subfield.map_le_iff_le_comap {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {f : K →+* L} {s : Subfield K} {t : Subfield L} :

                      range #

                      def RingHom.fieldRange {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) :

                      The range of a ring homomorphism, as a subfield of the target. See Note [range copy pattern].

                      Equations
                      Instances For
                        @[simp]
                        theorem RingHom.coe_fieldRange {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) :
                        @[simp]
                        theorem RingHom.mem_fieldRange {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {f : K →+* L} {y : L} :
                        y RingHom.fieldRange f ∃ (x : K), f x = y

                        The range of a morphism of fields is a fintype, if the domain is a fintype.

                        Note that this instance can cause a diamond with Subtype.Fintype if L is also a fintype.

                        Equations

                        inf #

                        The inf of two subfields is their intersection.

                        Equations
                        • Subfield.instInfSubfield = { inf := fun (s t : Subfield K) => let __src := s.toSubring t.toSubring; { toSubring := __src, inv_mem' := } }
                        @[simp]
                        theorem Subfield.coe_inf {K : Type u} [DivisionRing K] (p : Subfield K) (p' : Subfield K) :
                        (p p') = p.carrier p'.carrier
                        @[simp]
                        theorem Subfield.mem_inf {K : Type u} [DivisionRing K] {p : Subfield K} {p' : Subfield K} {x : K} :
                        x p p' x p x p'
                        Equations
                        • Subfield.instInfSetSubfield = { sInf := fun (S : Set (Subfield K)) => let __src := sInf (Subfield.toSubring '' S); { toSubring := __src, inv_mem' := } }
                        @[simp]
                        theorem Subfield.coe_sInf {K : Type u} [DivisionRing K] (S : Set (Subfield K)) :
                        (sInf S) = ⋂ s ∈ S, s
                        theorem Subfield.mem_sInf {K : Type u} [DivisionRing K] {S : Set (Subfield K)} {x : K} :
                        x sInf S pS, x p
                        @[simp]
                        theorem Subfield.sInf_toSubring {K : Type u} [DivisionRing K] (s : Set (Subfield K)) :
                        (sInf s).toSubring = ⨅ t ∈ s, t.toSubring
                        theorem Subfield.isGLB_sInf {K : Type u} [DivisionRing K] (S : Set (Subfield K)) :
                        IsGLB S (sInf S)

                        Subfields of a ring form a complete lattice.

                        Equations

                        subfield closure of a subset #

                        def Subfield.closure {K : Type u} [DivisionRing K] (s : Set K) :

                        The Subfield generated by a set.

                        Equations
                        Instances For
                          theorem Subfield.mem_closure {K : Type u} [DivisionRing K] {x : K} {s : Set K} :
                          x Subfield.closure s ∀ (S : Subfield K), s Sx S
                          @[simp]
                          theorem Subfield.subset_closure {K : Type u} [DivisionRing K] {s : Set K} :

                          The subfield generated by a set includes the set.

                          theorem Subfield.not_mem_of_not_mem_closure {K : Type u} [DivisionRing K] {s : Set K} {P : K} (hP : PSubfield.closure s) :
                          Ps
                          @[simp]
                          theorem Subfield.closure_le {K : Type u} [DivisionRing K] {s : Set K} {t : Subfield K} :

                          A subfield t includes closure s if and only if it includes s.

                          theorem Subfield.closure_mono {K : Type u} [DivisionRing K] ⦃s : Set K ⦃t : Set K (h : s t) :

                          Subfield closure of a set is monotone in its argument: if s ⊆ t, then closure s ≤ closure t.

                          theorem Subfield.closure_eq_of_le {K : Type u} [DivisionRing K] {s : Set K} {t : Subfield K} (h₁ : s t) (h₂ : t Subfield.closure s) :
                          theorem Subfield.closure_induction {K : Type u} [DivisionRing K] {s : Set K} {p : KProp} {x : K} (h : x Subfield.closure s) (mem : xs, p x) (one : p 1) (add : ∀ (x y : K), p xp yp (x + y)) (neg : ∀ (x : K), p xp (-x)) (inv : ∀ (x : K), p xp x⁻¹) (mul : ∀ (x y : K), p xp yp (x * y)) :
                          p x

                          An induction principle for closure membership. If p holds for 1, and all elements of s, and is preserved under addition, negation, and multiplication, then p holds for all elements of the closure of s.

                          def Subfield.gi (K : Type u) [DivisionRing K] :
                          GaloisInsertion Subfield.closure SetLike.coe

                          closure forms a Galois insertion with the coercion to set.

                          Equations
                          Instances For
                            theorem Subfield.closure_eq {K : Type u} [DivisionRing K] (s : Subfield K) :

                            Closure of a subfield S equals S.

                            @[simp]
                            theorem Subfield.closure_iUnion {K : Type u} [DivisionRing K] {ι : Sort u_1} (s : ιSet K) :
                            Subfield.closure (⋃ (i : ι), s i) = ⨆ (i : ι), Subfield.closure (s i)
                            theorem Subfield.map_sup {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (s : Subfield K) (t : Subfield K) (f : K →+* L) :
                            theorem Subfield.map_iSup {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {ι : Sort u_1} (f : K →+* L) (s : ιSubfield K) :
                            Subfield.map f (iSup s) = ⨆ (i : ι), Subfield.map f (s i)
                            theorem Subfield.comap_inf {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (s : Subfield L) (t : Subfield L) (f : K →+* L) :
                            theorem Subfield.comap_iInf {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] {ι : Sort u_1} (f : K →+* L) (s : ιSubfield L) :
                            Subfield.comap f (iInf s) = ⨅ (i : ι), Subfield.comap f (s i)
                            @[simp]
                            theorem Subfield.map_bot {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) :
                            @[simp]
                            theorem Subfield.comap_top {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) :
                            theorem Subfield.mem_iSup_of_directed {K : Type u} [DivisionRing K] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubfield K} (hS : Directed (fun (x x_1 : Subfield K) => x x_1) S) {x : K} :
                            x ⨆ (i : ι), S i ∃ (i : ι), x S i

                            The underlying set of a non-empty directed sSup of subfields is just a union of the subfields. Note that this fails without the directedness assumption (the union of two subfields is typically not a subfield)

                            theorem Subfield.coe_iSup_of_directed {K : Type u} [DivisionRing K] {ι : Sort u_1} [hι : Nonempty ι] {S : ιSubfield K} (hS : Directed (fun (x x_1 : Subfield K) => x x_1) S) :
                            (⨆ (i : ι), S i) = ⋃ (i : ι), (S i)
                            theorem Subfield.mem_sSup_of_directedOn {K : Type u} [DivisionRing K] {S : Set (Subfield K)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun (x x_1 : Subfield K) => x x_1) S) {x : K} :
                            x sSup S ∃ s ∈ S, x s
                            theorem Subfield.coe_sSup_of_directedOn {K : Type u} [DivisionRing K] {S : Set (Subfield K)} (Sne : Set.Nonempty S) (hS : DirectedOn (fun (x x_1 : Subfield K) => x x_1) S) :
                            (sSup S) = ⋃ s ∈ S, s

                            Restriction of a ring homomorphism to its range interpreted as a subfield.

                            Equations
                            Instances For
                              @[simp]
                              theorem RingHom.coe_rangeRestrictField {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) (x : K) :
                              def RingHom.eqLocusField {K : Type u} [DivisionRing K] {L : Type v} [Semiring L] (f : K →+* L) (g : K →+* L) :

                              The subfield of elements x : R such that f x = g x, i.e., the equalizer of f and g as a subfield of R

                              Equations
                              • One or more equations did not get rendered due to their size.
                              Instances For
                                theorem RingHom.eqOn_field_closure {K : Type u} [DivisionRing K] {L : Type v} [Semiring L] {f : K →+* L} {g : K →+* L} {s : Set K} (h : Set.EqOn (f) (g) s) :
                                Set.EqOn f g (Subfield.closure s)

                                If two ring homomorphisms are equal on a set, then they are equal on its subfield closure.

                                theorem RingHom.eq_of_eqOn_subfield_top {K : Type u} [DivisionRing K] {L : Type v} [Semiring L] {f : K →+* L} {g : K →+* L} (h : Set.EqOn f g ) :
                                f = g
                                theorem RingHom.eq_of_eqOn_of_field_closure_eq_top {K : Type u} [DivisionRing K] {L : Type v} [Semiring L] {s : Set K} (hs : Subfield.closure s = ) {f : K →+* L} {g : K →+* L} (h : Set.EqOn (f) (g) s) :
                                f = g
                                theorem RingHom.map_field_closure {K : Type u} {L : Type v} [DivisionRing K] [DivisionRing L] (f : K →+* L) (s : Set K) :

                                The image under a ring homomorphism of the subfield generated by a set equals the subfield generated by the image of the set.

                                def Subfield.inclusion {K : Type u} [DivisionRing K] {S : Subfield K} {T : Subfield K} (h : S T) :
                                S →+* T

                                The ring homomorphism associated to an inclusion of subfields.

                                Equations
                                Instances For
                                  def RingEquiv.subfieldCongr {K : Type u} [DivisionRing K] {s : Subfield K} {t : Subfield K} (h : s = t) :
                                  s ≃+* t

                                  Makes the identity isomorphism from a proof two subfields of a multiplicative monoid are equal.

                                  Equations
                                  Instances For
                                    theorem Subfield.multiset_prod_mem {K : Type u} [Field K] (s : Subfield K) (m : Multiset K) :
                                    (am, a s)Multiset.prod m s

                                    Product of a multiset of elements in a subfield is in the subfield.

                                    theorem Subfield.prod_mem {K : Type u} [Field K] (s : Subfield K) {ι : Type u_1} {t : Finset ι} {f : ιK} (h : ct, f c s) :
                                    (Finset.prod t fun (i : ι) => f i) s

                                    Product of elements of a subfield indexed by a Finset is in the subfield.

                                    theorem Subfield.mem_closure_iff {K : Type u} [Field K] {s : Set K} {x : K} :
                                    x Subfield.closure s ∃ y ∈ Subring.closure s, ∃ z ∈ Subring.closure s, y / z = x