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)
Subfield K
: the type of subfields of a division ringK
.
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
SubfieldClass S K
states S
is a type of subsets s ⊆ K
closed under field operations.
Instances
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
- ⋯ = ⋯
Equations
- SubfieldClass.instNNRatCast s = { nnratCast := fun (q : ℚ≥0) => ⟨↑q, ⋯⟩ }
Equations
- SubfieldClass.instRatCast s = { ratCast := fun (q : ℚ) => ⟨↑q, ⋯⟩ }
Alias of SubfieldClass.coe_ratCast
.
Alias of SubfieldClass.ratCast_mem
.
Alias of SubfieldClass.qsmul_mem
.
Equations
- SubfieldClass.instSMulNNRat s = { smul := fun (q : ℚ≥0) (x : ↥s) => ⟨q • ↑x, ⋯⟩ }
Equations
- SubfieldClass.instSMulRat s = { smul := fun (q : ℚ) (x : ↥s) => ⟨q • ↑x, ⋯⟩ }
A subfield inherits a division ring structure
Equations
- SubfieldClass.toDivisionRing S s = Function.Injective.divisionRing Subtype.val ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
A subfield of a field inherits a field structure
Equations
- SubfieldClass.toField S s = Function.Injective.field Subtype.val ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
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.
Stacks Tag 09FD (second part)
A subfield is closed under multiplicative inverses.
Instances For
The underlying AddSubgroup
of a subfield.
Equations
- s.toAddSubgroup = { toAddSubmonoid := s.toAddSubgroup.toAddSubmonoid, neg_mem' := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Two subfields are equal if they have the same elements.
Copy of a subfield with a new carrier
equal to the old one. Useful to fix definitional
equalities.
Equations
- S.copy s hs = { carrier := s, mul_mem' := ⋯, one_mem' := ⋯, add_mem' := ⋯, zero_mem' := ⋯, neg_mem' := ⋯, inv_mem' := ⋯ }
Instances For
A subfield contains the field's 1.
A subfield contains the field's 0.
A subfield is closed under multiplication.
A subfield is closed under addition.
A subfield is closed under negation.
A subfield is closed under subtraction.
A subfield is closed under inverses.
A subfield is closed under division.
Alias of intCast_mem
.
Equations
- s.instRingSubtypeMem = s.toRing
Equations
- s.toDivisionRing = Function.Injective.divisionRing Subtype.val ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
A subfield inherits a field structure
Equations
- s.toField = Function.Injective.field Subtype.val ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
The embedding from a subfield of the field K
to K
.
Equations
- s.subtype = { toFun := Subtype.val, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ }