Documentation

Mathlib.FieldTheory.Galois.Infinite

The Fundamental Theorem of Infinite Galois Theory #

In this file, we prove the fundamental theorem of infinite Galois theory and the special case for open subgroups and normal subgroups. We first verify that IntermediateField.fixingSubgroup and IntermediateField.fixedField are inverses of each other between intermediate fields and closed subgroups of the Galois group.

Main definitions and results #

In K/k, for any intermediate field L :

For any subgroup H of Gal(K/k) :

The fundamental theorem of infinite Galois theory :

Special cases :

theorem InfiniteGalois.fixingSubgroup_isClosed {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] (L : IntermediateField k K) [IsGalois k K] :
IsClosed L.fixingSubgroup
theorem InfiniteGalois.fixedField_fixingSubgroup {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] (L : IntermediateField k K) [IsGalois k K] :
IntermediateField.fixedField L.fixingSubgroup = L

For a subgroup H of Gal(K/k), the fixed field of the image of H under the restriction to a normal intermediate field E is equal to the fixed field of H in K intersecting with E.

theorem InfiniteGalois.fixingSubgroup_fixedField {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] (H : ClosedSubgroup (K ≃ₐ[k] K)) [IsGalois k K] :
(IntermediateField.fixedField H).fixingSubgroup = H

The Galois correspondence from intermediate fields to closed subgroups.

Equations
  • One or more equations did not get rendered due to their size.
Instances For
    def InfiniteGalois.GaloisInsertionIntermediateFieldClosedSubgroup {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] [IsGalois k K] :
    GaloisInsertion (OrderDual.toDual fun (E : IntermediateField k K) => { toSubgroup := E.fixingSubgroup, isClosed' := }) ((fun (H : ClosedSubgroup (K ≃ₐ[k] K)) => IntermediateField.fixedField H) OrderDual.toDual)

    The Galois correspondence as a GaloisInsertion

    Equations
    • InfiniteGalois.GaloisInsertionIntermediateFieldClosedSubgroup = InfiniteGalois.IntermediateFieldEquivClosedSubgroup.toGaloisInsertion
    Instances For
      def InfiniteGalois.GaloisCoinsertionIntermediateFieldSubgroup {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] [IsGalois k K] :
      GaloisCoinsertion (OrderDual.toDual fun (E : IntermediateField k K) => E.fixingSubgroup) ((fun (H : Subgroup (K ≃ₐ[k] K)) => IntermediateField.fixedField H) OrderDual.toDual)

      The Galois correspondence as a GaloisCoinsertion

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        theorem InfiniteGalois.isOpen_iff_finite {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] (L : IntermediateField k K) [IsGalois k K] :
        IsOpen (↑(InfiniteGalois.IntermediateFieldEquivClosedSubgroup L)).carrier FiniteDimensional k L
        theorem InfiniteGalois.normal_iff_isGalois {k : Type u_1} {K : Type u_2} [Field k] [Field K] [Algebra k K] (L : IntermediateField k K) [IsGalois k K] :
        (↑(InfiniteGalois.IntermediateFieldEquivClosedSubgroup L)).Normal IsGalois k L