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.top is the largest element of the series, and s.bot is the least element.

Two CompositionSeries X, s₁ and s₂ are equivalent if there is a bijection e : Fin s₁.lengthFin 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 both submodules and 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.

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) :

    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.top is the largest element of the series, and s.bot is the least element.

    Instances For
      theorem CompositionSeries.step' {X : Type u} [Lattice X] [JordanHolderLattice X] (self : CompositionSeries X) (i : Fin self.length) :
      JordanHolderLattice.IsMaximal (self.series i.castSucc) (self.series i.succ)
      instance CompositionSeries.coeFun {X : Type u} [Lattice X] [JordanHolderLattice X] :
      CoeFun (CompositionSeries X) fun (x : CompositionSeries X) => Fin (x.length + 1)X
      Equations
      • CompositionSeries.coeFun = { coe := CompositionSeries.series }
      Equations
      • CompositionSeries.inhabited = { default := { length := 0, series := default, step' := } }
      theorem CompositionSeries.step {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (i : Fin s.length) :
      JordanHolderLattice.IsMaximal (s.series i.castSucc) (s.series i.succ)
      theorem CompositionSeries.coeFn_mk {X : Type u} [Lattice X] [JordanHolderLattice X] (length : ) (series : Fin (length + 1)X) (step : ∀ (i : Fin length), JordanHolderLattice.IsMaximal (series i.castSucc) (series i.succ)) :
      { length := length, series := series, step' := step }.series = series
      theorem CompositionSeries.lt_succ {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (i : Fin s.length) :
      s.series i.castSucc < s.series 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.series i = s.series j i = j
      Equations
      theorem CompositionSeries.mem_def {X : Type u} [Lattice X] [JordanHolderLattice X] {x : X} {s : CompositionSeries X} :
      x s x Set.range s.series
      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

      The ordered List X of elements of a CompositionSeries X.

      Equations
      Instances For
        theorem CompositionSeries.ext_fun {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (hl : s₁.length = s₂.length) (h : ∀ (i : Fin (s₁.length + 1)), s₁.series i = s₂.series (Fin.cast i)) :
        s₁ = s₂

        Two CompositionSeries are equal if they are the same length and have the same ith element for every i

        @[simp]
        theorem CompositionSeries.length_toList {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
        s.toList.length = s.length + 1
        theorem CompositionSeries.chain'_toList {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
        List.Chain' JordanHolderLattice.IsMaximal s.toList
        theorem CompositionSeries.toList_sorted {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
        List.Sorted (fun (x x_1 : X) => x < x_1) s.toList
        @[simp]
        theorem CompositionSeries.mem_toList {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} :
        x s.toList x s
        def CompositionSeries.ofList {X : Type u} [Lattice X] [JordanHolderLattice X] (l : List X) (hl : l []) (hc : List.Chain' JordanHolderLattice.IsMaximal l) :

        Make a CompositionSeries X from the ordered list of its elements.

        Equations
        Instances For
          theorem CompositionSeries.length_ofList {X : Type u} [Lattice X] [JordanHolderLattice X] (l : List X) (hl : l []) (hc : List.Chain' JordanHolderLattice.IsMaximal l) :
          (CompositionSeries.ofList l hl hc).length = l.length - 1
          @[simp]
          theorem CompositionSeries.toList_ofList {X : Type u} [Lattice X] [JordanHolderLattice X] (l : List X) (hl : l []) (hc : List.Chain' JordanHolderLattice.IsMaximal l) :
          (CompositionSeries.ofList l hl hc).toList = l
          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.

          The largest element of a CompositionSeries

          Equations
          Instances For
            @[simp]
            theorem CompositionSeries.le_top {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (i : Fin (s.length + 1)) :
            s.series i s.top
            theorem CompositionSeries.le_top_of_mem {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} (hx : x s) :
            x s.top

            The smallest element of a CompositionSeries

            Equations
            • s.bot = s.series 0
            Instances For
              @[simp]
              theorem CompositionSeries.bot_le {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (i : Fin (s.length + 1)) :
              s.bot s.series i
              theorem CompositionSeries.bot_le_of_mem {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} (hx : x s) :
              s.bot x
              theorem CompositionSeries.length_pos_of_mem_ne {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} {y : X} (hx : x s) (hy : y s) (hxy : x y) :
              0 < s.length
              theorem CompositionSeries.forall_mem_eq_of_length_eq_zero {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (hs : s.length = 0) {x : X} {y : X} (hx : x s) (hy : y s) :
              x = y
              @[simp]
              theorem CompositionSeries.eraseTop_series {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (i : Fin (s.length - 1 + 1)) :
              s.eraseTop.series i = s.series i,
              @[simp]
              theorem CompositionSeries.eraseTop_length {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
              s.eraseTop.length = s.length - 1

              Remove the largest element from a CompositionSeries. If the series s has length zero, then s.eraseTop = s

              Equations
              • s.eraseTop = { length := s.length - 1, series := fun (i : Fin (s.length - 1 + 1)) => s.series i, , step' := }
              Instances For
                theorem CompositionSeries.top_eraseTop {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
                s.eraseTop.top = s.series s.length - 1,
                @[simp]
                theorem CompositionSeries.bot_eraseTop {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) :
                s.eraseTop.bot = s.bot
                theorem CompositionSeries.mem_eraseTop_of_ne_of_mem {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} (hx : x s.top) (hxs : x s) :
                x s.eraseTop
                theorem CompositionSeries.mem_eraseTop {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} (h : 0 < s.length) :
                x s.eraseTop x s.top x s
                theorem CompositionSeries.lt_top_of_mem_eraseTop {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} (h : 0 < s.length) (hx : x s.eraseTop) :
                x < s.top
                theorem CompositionSeries.append_castAdd_aux {α : Type u_1} {m : } {n : } (a : Fin m.succα) (b : Fin n.succα) (i : Fin m) :
                Matrix.vecAppend (a Fin.castSucc) b (Fin.castAdd n i).castSucc = a i.castSucc
                theorem CompositionSeries.append_succ_castAdd_aux {α : Type u_1} {m : } {n : } (a : Fin m.succα) (b : Fin n.succα) (i : Fin m) (h : a (Fin.last m) = b 0) :
                Matrix.vecAppend (a Fin.castSucc) b (Fin.castAdd n i).succ = a i.succ
                theorem CompositionSeries.append_natAdd_aux {α : Type u_1} {m : } {n : } (a : Fin m.succα) (b : Fin n.succα) (i : Fin n) :
                Matrix.vecAppend (a Fin.castSucc) b (Fin.natAdd m i).castSucc = b i.castSucc
                theorem CompositionSeries.append_succ_natAdd_aux {α : Type u_1} {m : } {n : } (a : Fin m.succα) (b : Fin n.succα) (i : Fin n) :
                Matrix.vecAppend (a Fin.castSucc) b (Fin.natAdd m i).succ = b i.succ
                @[simp]
                theorem CompositionSeries.append_length {X : Type u} [Lattice X] [JordanHolderLattice X] (s₁ : CompositionSeries X) (s₂ : CompositionSeries X) (h : s₁.top = s₂.bot) :
                (s₁.append s₂ h).length = s₁.length + s₂.length
                def CompositionSeries.append {X : Type u} [Lattice X] [JordanHolderLattice X] (s₁ : CompositionSeries X) (s₂ : CompositionSeries X) (h : s₁.top = s₂.bot) :

                Append two composition series s₁ and s₂ such that the least element of s₁ is the maximum element of s₂.

                Equations
                • s₁.append s₂ h = { length := s₁.length + s₂.length, series := Matrix.vecAppend (s₁.series Fin.castSucc) s₂.series, step' := }
                Instances For
                  theorem CompositionSeries.coe_append {X : Type u} [Lattice X] [JordanHolderLattice X] (s₁ : CompositionSeries X) (s₂ : CompositionSeries X) (h : s₁.top = s₂.bot) :
                  (s₁.append s₂ h).series = Matrix.vecAppend (s₁.series Fin.castSucc) s₂.series
                  @[simp]
                  theorem CompositionSeries.append_castAdd {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.top = s₂.bot) (i : Fin s₁.length) :
                  (s₁.append s₂ h).series (Fin.castAdd s₂.length i).castSucc = s₁.series i.castSucc
                  @[simp]
                  theorem CompositionSeries.append_succ_castAdd {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.top = s₂.bot) (i : Fin s₁.length) :
                  (s₁.append s₂ h).series (Fin.castAdd s₂.length i).succ = s₁.series i.succ
                  @[simp]
                  theorem CompositionSeries.append_natAdd {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.top = s₂.bot) (i : Fin s₂.length) :
                  (s₁.append s₂ h).series (Fin.natAdd s₁.length i).castSucc = s₂.series i.castSucc
                  @[simp]
                  theorem CompositionSeries.append_succ_natAdd {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (h : s₁.top = s₂.bot) (i : Fin s₂.length) :
                  (s₁.append s₂ h).series (Fin.natAdd s₁.length i).succ = s₂.series i.succ
                  @[simp]
                  theorem CompositionSeries.snoc_length {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hsat : JordanHolderLattice.IsMaximal s.top x) :
                  (s.snoc x hsat).length = s.length + 1

                  Add an element to the top of a CompositionSeries

                  Equations
                  • s.snoc x hsat = { length := s.length + 1, series := Fin.snoc s.series x, step' := }
                  Instances For
                    @[simp]
                    theorem CompositionSeries.top_snoc {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hsat : JordanHolderLattice.IsMaximal s.top x) :
                    (s.snoc x hsat).top = x
                    @[simp]
                    theorem CompositionSeries.snoc_last {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hsat : JordanHolderLattice.IsMaximal s.top x) :
                    (s.snoc x hsat).series (Fin.last (s.length + 1)) = x
                    @[simp]
                    theorem CompositionSeries.snoc_castSucc {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hsat : JordanHolderLattice.IsMaximal s.top x) (i : Fin (s.length + 1)) :
                    (s.snoc x hsat).series i.castSucc = s.series i
                    @[simp]
                    theorem CompositionSeries.bot_snoc {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hsat : JordanHolderLattice.IsMaximal s.top x) :
                    (s.snoc x hsat).bot = s.bot
                    theorem CompositionSeries.mem_snoc {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} {x : X} {y : X} {hsat : JordanHolderLattice.IsMaximal s.top x} :
                    y s.snoc x hsat y s y = x
                    theorem CompositionSeries.eq_snoc_eraseTop {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (h : 0 < s.length) :
                    s = s.eraseTop.snoc s.top
                    @[simp]
                    theorem CompositionSeries.snoc_eraseTop_top {X : Type u} [Lattice X] [JordanHolderLattice X] {s : CompositionSeries X} (h : JordanHolderLattice.IsMaximal s.eraseTop.top s.top) :
                    s.eraseTop.snoc s.top h = s

                    Two CompositionSeries X, s₁ and s₂ are equivalent if there is a bijection e : Fin s₁.lengthFin 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.append {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} {t₁ : CompositionSeries X} {t₂ : CompositionSeries X} (hs : s₁.top = s₂.bot) (ht : t₁.top = t₂.bot) (h₁ : s₁.Equivalent t₁) (h₂ : s₂.Equivalent t₂) :
                      (s₁.append s₂ hs).Equivalent (t₁.append t₂ ht)
                      theorem CompositionSeries.Equivalent.snoc {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} {x₁ : X} {x₂ : X} {hsat₁ : JordanHolderLattice.IsMaximal s₁.top x₁} {hsat₂ : JordanHolderLattice.IsMaximal s₂.top x₂} (hequiv : s₁.Equivalent s₂) (htop : JordanHolderLattice.Iso (s₁.top, x₁) (s₂.top, x₂)) :
                      (s₁.snoc x₁ hsat₁).Equivalent (s₂.snoc 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 s.top x₁} {hsat₂ : JordanHolderLattice.IsMaximal s.top x₂} {hsaty₁ : JordanHolderLattice.IsMaximal (s.snoc x₁ hsat₁).top y₁} {hsaty₂ : JordanHolderLattice.IsMaximal (s.snoc x₂ hsat₂).top y₂} (hr₁ : JordanHolderLattice.Iso (s.top, x₁) (x₂, y₂)) (hr₂ : JordanHolderLattice.Iso (x₁, y₁) (s.top, x₂)) :
                      ((s.snoc x₁ hsat₁).snoc y₁ hsaty₁).Equivalent ((s.snoc x₂ hsat₂).snoc y₂ hsaty₂)
                      theorem CompositionSeries.length_eq_zero_of_bot_eq_bot_of_top_eq_top_of_length_eq_zero {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (hb : s₁.bot = s₂.bot) (ht : s₁.top = s₂.top) (hs₁ : s₁.length = 0) :
                      s₂.length = 0
                      theorem CompositionSeries.length_pos_of_bot_eq_bot_of_top_eq_top_of_length_pos {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (hb : s₁.bot = s₂.bot) (ht : s₁.top = s₂.top) :
                      0 < s₁.length0 < s₂.length
                      theorem CompositionSeries.eq_of_bot_eq_bot_of_top_eq_top_of_length_eq_zero {X : Type u} [Lattice X] [JordanHolderLattice X] {s₁ : CompositionSeries X} {s₂ : CompositionSeries X} (hb : s₁.bot = s₂.bot) (ht : s₁.top = s₂.top) (hs₁0 : s₁.length = 0) :
                      s₁ = s₂
                      theorem CompositionSeries.exists_top_eq_snoc_equivalant {X : Type u} [Lattice X] [JordanHolderLattice X] (s : CompositionSeries X) (x : X) (hm : JordanHolderLattice.IsMaximal x s.top) (hb : s.bot x) :
                      ∃ (t : CompositionSeries X), t.bot = s.bot t.length + 1 = s.length ∃ (htx : t.top = x), s.Equivalent (t.snoc s.top )

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

                      theorem CompositionSeries.jordan_holder {X : Type u} [Lattice X] [JordanHolderLattice X] (s₁ : CompositionSeries X) (s₂ : CompositionSeries X) (hb : s₁.bot = s₂.bot) (ht : s₁.top = s₂.top) :
                      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.