Documentation

Mathlib.Order.JordanHolder

Jordan-Hölder Theorem #

This file proves the Jordan Hölder theorem for a JordanHolderLattice, a class also defined in this file. Examples of JordanHolderLattice include Subgroup G if G is a group, and Submodule R M if M is an R-module. Using this approach the theorem need not be proved separately for both groups and modules, the proof in this file can be applied to both.

Main definitions #

The main definitions in this file are JordanHolderLattice and CompositionSeries, and the relation Equivalent on CompositionSeries

A JordanHolderLattice is the class for which the Jordan Hölder theorem is proved. A Jordan Hölder lattice is a lattice equipped with a notion of maximality, IsMaximal, and a notion of isomorphism of pairs Iso. In the example of subgroups of a group, IsMaximal H K means that H is a maximal normal subgroup of K, and Iso (H₁, K₁) (H₂, K₂) means that the quotient H₁ / K₁ is isomorphic to the quotient H₂ / K₂. Iso must be symmetric and transitive and must satisfy the second isomorphism theorem Iso (H, H ⊔ K) (H ⊓ K, K).

A CompositionSeries X is a finite nonempty series of elements of the lattice X such that each element is maximal inside the next. The length of a CompositionSeries X is one less than the number of elements in the series. Note that there is no stipulation that a series start from the bottom of the lattice and finish at the top. For a composition series s, s.last is the largest element of the series, and s.head is the least element.

Two CompositionSeries X, s₁ and s₂ are equivalent if there is a bijection e : Fin s₁.length ≃ Fin s₂.length such that for any i, Iso (s₁ i, s₁ i.succ) (s₂ (e i), s₂ (e i.succ))

Main theorems #

The main theorem is CompositionSeries.jordan_holder, which says that if two composition series have the same least element and the same largest element, then they are Equivalent.

TODO #

Provide instances of JordanHolderLattice for subgroups, and potentially for modular lattices.

It is not entirely clear how this should be done. Possibly there should be no global instances of JordanHolderLattice, and the instances should only be defined locally in order to prove the Jordan-Hölder theorem for modules/groups and the API should be transferred because many of the theorems in this file will have stronger versions for modules. There will also need to be an API for mapping composition series across homomorphisms. It is also probably possible to provide an instance of JordanHolderLattice for any ModularLattice, and in this case the Jordan-Hölder theorem will say that there is a well defined notion of length of a modular lattice. However an instance of JordanHolderLattice for a modular lattice will not be able to contain the correct notion of isomorphism for modules, so a separate instance for modules will still be required and this will clash with the instance for modular lattices, and so at least one of these instances should not be a global instance.

[!NOTE] The previous paragraph indicates that the instance of JordanHolderLattice for submodules should be obtained via ModularLattice. This is not the case in mathlib4. See JordanHolderModule.instJordanHolderLattice.

class JordanHolderLattice (X : Type u) [Lattice X] :

A JordanHolderLattice is the class for which the Jordan Hölder theorem is proved. A Jordan Hölder lattice is a lattice equipped with a notion of maximality, IsMaximal, and a notion of isomorphism of pairs Iso. In the example of subgroups of a group, IsMaximal H K means that H is a maximal normal subgroup of K, and Iso (H₁, K₁) (H₂, K₂) means that the quotient H₁ / K₁ is isomorphic to the quotient H₂ / K₂. Iso must be symmetric and transitive and must satisfy the second isomorphism theorem Iso (H, H ⊔ K) (H ⊓ K, K). Examples include Subgroup G if G is a group, and Submodule R M if M is an R-module.

Instances
    theorem JordanHolderLattice.sup_eq_of_isMaximal {X : Type u} [Lattice X] [self : JordanHolderLattice X] {x : X} {y : X} {z : X} :
    theorem JordanHolderLattice.second_iso {X : Type u} [Lattice X] [self : JordanHolderLattice X] {x : X} {y : X} :
    theorem JordanHolderLattice.isMaximal_of_eq_inf {X : Type u} [Lattice X] [JordanHolderLattice X] (x : X) (b : X) {a : X} {y : X} (ha : x y = a) (hxy : x y) (hxb : JordanHolderLattice.IsMaximal x b) (hyb : JordanHolderLattice.IsMaximal y b) :
    theorem JordanHolderLattice.second_iso_of_eq {X : Type u} [Lattice X] [JordanHolderLattice X] {x : X} {y : X} {a : X} {b : X} (hm : JordanHolderLattice.IsMaximal x a) (ha : x y = a) (hb : x y = b) :
    @[reducible, inline]

    A CompositionSeries X is a finite nonempty series of elements of a JordanHolderLattice such that each element is maximal inside the next. The length of a CompositionSeries X is one less than the number of elements in the series. Note that there is no stipulation that a series start from the bottom of the lattice and finish at the top. For a composition series s, s.last is the largest element of the series, and s.head is the least element.

    Equations
    Instances For
      theorem CompositionSeries.lt_succ {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (i : Fin s.length) :
      s.toFun i.castSucc < s.toFun i.succ
      @[simp]
      theorem CompositionSeries.inj {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) {i : Fin s.length.succ} {j : Fin s.length.succ} :
      s.toFun i = s.toFun j i = j
      theorem CompositionSeries.total {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} {y : X} (hx : x s) (hy : y s) :
      x y y x
      theorem CompositionSeries.ext {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : ∀ (x : X), x s₁ x s₂) :
      s₁ = s₂

      Two CompositionSeries are equal if they have the same elements. See also ext_fun.

      @[simp]
      theorem CompositionSeries.le_last {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (i : Fin (s.length + 1)) :
      s.toFun i RelSeries.last s
      @[simp]
      theorem CompositionSeries.head_le {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (i : Fin (s.length + 1)) :
      RelSeries.head s s.toFun i

      Two CompositionSeries X, s₁ and s₂ are equivalent if there is a bijection e : Fin s₁.length ≃ Fin s₂.length such that for any i, Iso (s₁ i) (s₁ i.succ) (s₂ (e i), s₂ (e i.succ))

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For
        theorem CompositionSeries.Equivalent.symm {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.Equivalent s₂) :
        s₂.Equivalent s₁
        theorem CompositionSeries.Equivalent.trans {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} {s₃ : CompositionSeries X} (h₁ : s₁.Equivalent s₂) (h₂ : s₂.Equivalent s₃) :
        s₁.Equivalent s₃
        theorem CompositionSeries.Equivalent.smash {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} {t₁ : CompositionSeries X} {t₂ : CompositionSeries X} (hs : RelSeries.last s₁ = RelSeries.head s₂) (ht : RelSeries.last t₁ = RelSeries.head t₂) (h₁ : s₁.Equivalent t₁) (h₂ : s₂.Equivalent t₂) :
        theorem CompositionSeries.Equivalent.snoc {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} {x₁ : X} {x₂ : X} {hsat₁ : JordanHolderLattice.IsMaximal (RelSeries.last s₁) x₁} {hsat₂ : JordanHolderLattice.IsMaximal (RelSeries.last s₂) x₂} (hequiv : s₁.Equivalent s₂) (hlast : JordanHolderLattice.Iso (RelSeries.last s₁, x₁) (RelSeries.last s₂, x₂)) :
        CompositionSeries.Equivalent (RelSeries.snoc s₁ x₁ hsat₁) (RelSeries.snoc s₂ x₂ hsat₂)
        theorem CompositionSeries.Equivalent.length_eq {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.Equivalent s₂) :
        s₁.length = s₂.length
        theorem CompositionSeries.Equivalent.snoc_snoc_swap {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x₁ : X} {x₂ : X} {y₁ : X} {y₂ : X} {hsat₁ : JordanHolderLattice.IsMaximal (RelSeries.last s) x₁} {hsat₂ : JordanHolderLattice.IsMaximal (RelSeries.last s) x₂} {hsaty₁ : JordanHolderLattice.IsMaximal (RelSeries.snoc s x₁ hsat₁).last y₁} {hsaty₂ : JordanHolderLattice.IsMaximal (RelSeries.snoc s x₂ hsat₂).last y₂} (hr₁ : JordanHolderLattice.Iso (RelSeries.last s, x₁) (x₂, y₂)) (hr₂ : JordanHolderLattice.Iso (x₁, y₁) (RelSeries.last s, x₂)) :
        CompositionSeries.Equivalent ((RelSeries.snoc s x₁ hsat₁).snoc y₁ hsaty₁) ((RelSeries.snoc s x₂ hsat₂).snoc y₂ hsaty₂)

        Given a CompositionSeries, s, and an element x such that x is maximal inside s.last there is a series, t, such that t.last = x, t.head = s.head and snoc t s.last _ is equivalent to s.

        theorem CompositionSeries.jordan_holder {X : Type u} [Lattice X] [JordanHolderLattice X] (s₁ : CompositionSeries X) (s₂ : CompositionSeries X) (hb : RelSeries.head s₁ = RelSeries.head s₂) (ht : RelSeries.last s₁ = RelSeries.last s₂) :
        s₁.Equivalent s₂

        The Jordan-Hölder theorem, stated for any JordanHolderLattice. If two composition series start and finish at the same place, they are equivalent.