Jordan-Hölder Theorem #
THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.
This file proves the Jordan Hölder theorem for a jordan_holder_lattice, a class also defined in
this file. Examples of jordan_holder_lattice 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
seperately for both groups and modules, the proof in this file can be applied to both.
Main definitions #
The main definitions in this file are jordan_holder_lattice and composition_series,
and the relation equivalent on composition_series
A jordan_holder_lattice 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, is_maximal, and a notion
of isomorphism of pairs iso. In the example of subgroups of a group, is_maximal 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 composition_series 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 composition_series 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 composition_series 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 composition_series.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 jordan_holder_lattice 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 jordan_holder_lattice, 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 jordan_holder_lattice for any modular_lattice, 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 jordan_holder_lattice 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.
- is_maximal : X → X → Prop
- lt_of_is_maximal : ∀ {x y : X}, jordan_holder_lattice.is_maximal x y → x < y
- sup_eq_of_is_maximal : ∀ {x y z : X}, jordan_holder_lattice.is_maximal x z → jordan_holder_lattice.is_maximal y z → x ≠ y → x ⊔ y = z
- is_maximal_inf_left_of_is_maximal_sup : ∀ {x y : X}, jordan_holder_lattice.is_maximal x (x ⊔ y) → jordan_holder_lattice.is_maximal y (x ⊔ y) → jordan_holder_lattice.is_maximal (x ⊓ y) x
- iso : X × X → X × X → Prop
- iso_symm : ∀ {x y : X × X}, jordan_holder_lattice.iso x y → jordan_holder_lattice.iso y x
- iso_trans : ∀ {x y z : X × X}, jordan_holder_lattice.iso x y → jordan_holder_lattice.iso y z → jordan_holder_lattice.iso x z
- second_iso : ∀ {x y : X}, jordan_holder_lattice.is_maximal x (x ⊔ y) → jordan_holder_lattice.iso (x, x ⊔ y) (x ⊓ y, y)
A jordan_holder_lattice 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, is_maximal, and a notion
of isomorphism of pairs iso. In the example of subgroups of a group, is_maximal 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 of this typeclass
Instances of other typeclasses for jordan_holder_lattice
        
        - jordan_holder_lattice.has_sizeof_inst
- length : ℕ
- series : fin (self.length + 1) → X
- step' : ∀ (i : fin self.length), jordan_holder_lattice.is_maximal (self.series (⇑fin.cast_succ i)) (self.series i.succ)
A composition_series X is a finite nonempty series of elements of a
jordan_holder_lattice such that each element is maximal inside the next. The length of a
composition_series 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 composition_series
        
        - composition_series.has_sizeof_inst
- composition_series.has_coe_to_fun
- composition_series.inhabited
- composition_series.has_mem
Equations
Equations
- composition_series.inhabited = {default := {length := 0, series := inhabited.default (pi.inhabited (fin (0 + 1))), step' := _}}
Equations
- composition_series.has_mem = {mem := λ (x : X) (s : composition_series X), x ∈ set.range ⇑s}
The ordered list X of elements of a composition_series X.
Equations
- s.to_list = list.of_fn ⇑s
Two composition_series are equal if they are the same length and
have the same ith element for every i
Make a composition_series X from the ordered list of its elements.
Two composition_series are equal if they have the same elements. See also ext_fun.
The largest element of a composition_series
The smallest element of a composition_series
Remove the largest element from a composition_series. If the series s
has length zero, then s.erase_top = s
Append two composition series s₁ and s₂ such that
the least element of s₁ is the maximum element of s₂.
Add an element to the top of a composition_series
Two composition_series 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))
Given a composition_series, 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.
The Jordan-Hölder theorem, stated for any jordan_holder_lattice.
If two composition series start and finish at the same place, they are equivalent.