aboutsummaryrefslogtreecommitdiff
path: root/mathcomp/solvable
diff options
context:
space:
mode:
authorEnrico Tassi2015-03-09 11:07:53 +0100
committerEnrico Tassi2015-03-09 11:24:38 +0100
commitfc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch)
treec16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/solvable
Initial commit
Diffstat (limited to 'mathcomp/solvable')
-rw-r--r--mathcomp/solvable/abelian.v2161
-rw-r--r--mathcomp/solvable/all.v19
-rw-r--r--mathcomp/solvable/alt.v528
-rw-r--r--mathcomp/solvable/burnside_app.v1305
-rw-r--r--mathcomp/solvable/center.v652
-rw-r--r--mathcomp/solvable/commutator.v362
-rw-r--r--mathcomp/solvable/extraspecial.v833
-rw-r--r--mathcomp/solvable/extremal.v2331
-rw-r--r--mathcomp/solvable/finmodule.v596
-rw-r--r--mathcomp/solvable/frobenius.v794
-rw-r--r--mathcomp/solvable/gfunctor.v484
-rw-r--r--mathcomp/solvable/gseries.v546
-rw-r--r--mathcomp/solvable/hall.v895
-rw-r--r--mathcomp/solvable/jordanholder.v681
-rw-r--r--mathcomp/solvable/maximal.v1656
-rw-r--r--mathcomp/solvable/nilpotent.v755
-rw-r--r--mathcomp/solvable/pgroup.v1355
-rw-r--r--mathcomp/solvable/primitive_action.v347
-rw-r--r--mathcomp/solvable/sylow.v661
-rw-r--r--mathcomp/solvable/wielandt_fixpoint.v651
20 files changed, 17612 insertions, 0 deletions
diff --git a/mathcomp/solvable/abelian.v b/mathcomp/solvable/abelian.v
new file mode 100644
index 0000000..8c412dc
--- /dev/null
+++ b/mathcomp/solvable/abelian.v
@@ -0,0 +1,2161 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype.
+Require Import finfun bigop finset prime binomial fingroup morphism perm.
+Require Import automorphism action quotient gfunctor gproduct zmodp cyclic.
+Require Import pgroup gseries nilpotent sylow.
+
+(******************************************************************************)
+(* Constructions based on abelian groups and their structure, with some *)
+(* emphasis on elementary abelian p-groups. *)
+(* 'Ldiv_n() == the set of all x that satisfy x ^+ n = 1, or, *)
+(* equivalently the set of x whose order divides n. *)
+(* 'Ldiv_n(G) == the set of x in G that satisfy x ^+ n = 1. *)
+(* := G :&: 'Ldiv_n() (pure Notation) *)
+(* exponent G == the exponent of G: the least e such that x ^+ e = 1 *)
+(* for all x in G (the LCM of the orders of x \in G). *)
+(* If G is nilpotent its exponent is reached. Note that *)
+(* `exponent G %| m' reads as `G has exponent m'. *)
+(* 'm(G) == the generator rank of G: the size of a smallest *)
+(* generating set for G (this is a basis for G if G *)
+(* abelian). *)
+(* abelian_type G == the abelian type of G : if G is abelian, a lexico- *)
+(* graphically maximal sequence of the orders of the *)
+(* elements of a minimal basis of G (if G is a p-group *)
+(* this is the sequence of orders for any basis of G, *)
+(* sorted in decending order). *)
+(* homocyclic G == G is the direct product of cycles of equal order, *)
+(* i.e., G is abelian with constant abelian type. *)
+(* p.-abelem G == G is an elementary abelian p-group, i.e., it is *)
+(* an abelian p-group of exponent p, and thus of order *)
+(* p ^ 'm(G) and rank (logn p #|G|). *)
+(* is_abelem G == G is an elementary abelian p-group for some prime p. *)
+(* 'E_p(G) == the set of elementary abelian p-subgroups of G. *)
+(* := [set E : {group _} | p.-abelem E & E \subset G] *)
+(* 'E_p^n(G) == the set of elementary abelian p-subgroups of G of *)
+(* order p ^ n (or, equivalently, of rank n). *)
+(* := [set E in 'E_p(G) | logn p #|E| == n] *)
+(* := [set E in 'E_p(G) | #|E| == p ^ n]%N if p is prime *)
+(* 'E*_p(G) == the set of maximal elementary abelian p-subgroups *)
+(* of G. *)
+(* := [set E | [max E | E \in 'E_p(G)]] *)
+(* 'E^n(G) == the set of elementary abelian subgroups of G that *)
+(* have gerank n (i.e., p-rank n for some prime p). *)
+(* := \bigcup_(0 <= p < #|G|.+1) 'E_p^n(G) *)
+(* 'r_p(G) == the p-rank of G: the maximal rank of an elementary *)
+(* subgroup of G. *)
+(* := \max_(E in 'E_p(G)) logn p #|E|. *)
+(* 'r(G) == the rank of G. *)
+(* := \max_(0 <= p < #|G|.+1) 'm_p(G). *)
+(* Note that 'r(G) coincides with 'r_p(G) if G is a p-group, and with 'm(G) *)
+(* if G is abelian, but is much more useful than 'm(G) in the proof of the *)
+(* Odd Order Theorem. *)
+(* 'Ohm_n(G) == the group generated by the x in G with order p ^ m *)
+(* for some prime p and some m <= n. Usually, G will be *)
+(* a p-group, so 'Ohm_n(G) will be generated by *)
+(* 'Ldiv_(p ^ n)(G), set of elements of G of order at *)
+(* most p ^ n. If G is also abelian then 'Ohm_n(G) *)
+(* consists exactly of those element, and the abelian *)
+(* type of G can be computed from the orders of the *)
+(* 'Ohm_n(G) subgroups. *)
+(* 'Mho^n(G) == the group generated by the x ^+ (p ^ n) for x a *)
+(* p-element of G for some prime p. Usually G is a *)
+(* p-group, and 'Mho^n(G) is generated by all such *)
+(* x ^+ (p ^ n); it consists of exactly these if G is *)
+(* also abelian. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section AbelianDefs.
+
+(* We defer the definition of the functors ('Omh_n(G), 'Mho^n(G)) because *)
+(* they must quantify over the finGroupType explicitly. *)
+
+Variable gT : finGroupType.
+Implicit Types (x : gT) (A B : {set gT}) (pi : nat_pred) (p n : nat).
+
+Definition Ldiv n := [set x : gT | x ^+ n == 1].
+
+Definition exponent A := \big[lcmn/1%N]_(x in A) #[x].
+
+Definition abelem p A := [&& p.-group A, abelian A & exponent A %| p].
+
+Definition is_abelem A := abelem (pdiv #|A|) A.
+
+Definition pElem p A := [set E : {group gT} | E \subset A & abelem p E].
+
+Definition pnElem p n A := [set E in pElem p A | logn p #|E| == n].
+
+Definition nElem n A := \bigcup_(0 <= p < #|A|.+1) pnElem p n A.
+
+Definition pmaxElem p A := [set E | [max E | E \in pElem p A]].
+
+Definition p_rank p A := \max_(E in pElem p A) logn p #|E|.
+
+Definition rank A := \max_(0 <= p < #|A|.+1) p_rank p A.
+
+Definition gen_rank A := #|[arg min_(B < A | <<B>> == A) #|B|]|.
+
+(* The definition of abelian_type depends on an existence lemma. *)
+(* The definition of homocyclic depends on abelian_type. *)
+
+End AbelianDefs.
+
+Arguments Scope exponent [_ group_scope].
+Arguments Scope abelem [_ nat_scope group_scope].
+Arguments Scope is_abelem [_ group_scope].
+Arguments Scope pElem [_ nat_scope group_scope].
+Arguments Scope pnElem [_ nat_scope nat_scope group_scope].
+Arguments Scope nElem [_ nat_scope group_scope].
+Arguments Scope pmaxElem [_ nat_scope group_scope].
+Arguments Scope p_rank [_ nat_scope group_scope].
+Arguments Scope rank [_ group_scope].
+Arguments Scope gen_rank [_ group_scope].
+
+Notation "''Ldiv_' n ()" := (Ldiv _ n)
+ (at level 8, n at level 2, format "''Ldiv_' n ()") : group_scope.
+
+Notation "''Ldiv_' n ( G )" := (G :&: 'Ldiv_n())
+ (at level 8, n at level 2, format "''Ldiv_' n ( G )") : group_scope.
+
+Prenex Implicits exponent.
+
+Notation "p .-abelem" := (abelem p)
+ (at level 2, format "p .-abelem") : group_scope.
+
+Notation "''E_' p ( G )" := (pElem p G)
+ (at level 8, p at level 2, format "''E_' p ( G )") : group_scope.
+
+Notation "''E_' p ^ n ( G )" := (pnElem p n G)
+ (at level 8, p, n at level 2, format "''E_' p ^ n ( G )") : group_scope.
+
+Notation "''E' ^ n ( G )" := (nElem n G)
+ (at level 8, n at level 2, format "''E' ^ n ( G )") : group_scope.
+
+Notation "''E*_' p ( G )" := (pmaxElem p G)
+ (at level 8, p at level 2, format "''E*_' p ( G )") : group_scope.
+
+Notation "''m' ( A )" := (gen_rank A)
+ (at level 8, format "''m' ( A )") : group_scope.
+
+Notation "''r' ( A )" := (rank A)
+ (at level 8, format "''r' ( A )") : group_scope.
+
+Notation "''r_' p ( A )" := (p_rank p A)
+ (at level 8, p at level 2, format "''r_' p ( A )") : group_scope.
+
+Section Functors.
+
+(* A functor needs to quantify over the finGroupType just beore the set. *)
+
+Variables (n : nat) (gT : finGroupType) (A : {set gT}).
+
+Definition Ohm := <<[set x in A | x ^+ (pdiv #[x] ^ n) == 1]>>.
+
+Definition Mho := <<[set x ^+ (pdiv #[x] ^ n) | x in A & (pdiv #[x]).-elt x]>>.
+
+Canonical Ohm_group : {group gT} := Eval hnf in [group of Ohm].
+Canonical Mho_group : {group gT} := Eval hnf in [group of Mho].
+
+Lemma pdiv_p_elt (p : nat) (x : gT) : p.-elt x -> x != 1 -> pdiv #[x] = p.
+Proof.
+move=> p_x; rewrite /order -cycle_eq1.
+by case/(pgroup_pdiv p_x)=> p_pr _ [k ->]; rewrite pdiv_pfactor.
+Qed.
+
+Lemma OhmPredP (x : gT) :
+ reflect (exists2 p, prime p & x ^+ (p ^ n) = 1) (x ^+ (pdiv #[x] ^ n) == 1).
+Proof.
+have [-> | nt_x] := eqVneq x 1.
+ by rewrite expg1n eqxx; left; exists 2; rewrite ?expg1n.
+apply: (iffP idP) => [/eqP | [p p_pr /eqP x_pn]].
+ by exists (pdiv #[x]); rewrite ?pdiv_prime ?order_gt1.
+rewrite (@pdiv_p_elt p) //; rewrite -order_dvdn in x_pn.
+by rewrite [p_elt _ _](pnat_dvd x_pn) // pnat_exp pnat_id.
+Qed.
+
+Lemma Mho_p_elt (p : nat) x : x \in A -> p.-elt x -> x ^+ (p ^ n) \in Mho.
+Proof.
+move=> Ax p_x; case: (eqVneq x 1) => [-> | ntx]; first by rewrite groupX.
+by apply: mem_gen; apply/imsetP; exists x; rewrite ?inE ?Ax (pdiv_p_elt p_x).
+Qed.
+
+End Functors.
+
+Arguments Scope Ohm [nat_scope _ group_scope].
+Arguments Scope Ohm_group [nat_scope _ group_scope].
+Arguments Scope Mho [nat_scope _ group_scope].
+Arguments Scope Mho_group [nat_scope _ group_scope].
+Implicit Arguments OhmPredP [n gT x].
+
+Notation "''Ohm_' n ( G )" := (Ohm n G)
+ (at level 8, n at level 2, format "''Ohm_' n ( G )") : group_scope.
+Notation "''Ohm_' n ( G )" := (Ohm_group n G) : Group_scope.
+
+Notation "''Mho^' n ( G )" := (Mho n G)
+ (at level 8, n at level 2, format "''Mho^' n ( G )") : group_scope.
+Notation "''Mho^' n ( G )" := (Mho_group n G) : Group_scope.
+
+Section ExponentAbelem.
+
+Variable gT : finGroupType.
+Implicit Types (p n : nat) (pi : nat_pred) (x : gT) (A B C : {set gT}).
+Implicit Types E G H K P X Y : {group gT}.
+
+Lemma LdivP A n x : reflect (x \in A /\ x ^+ n = 1) (x \in 'Ldiv_n(A)).
+Proof. by rewrite !inE; apply: (iffP andP) => [] [-> /eqP]. Qed.
+
+Lemma dvdn_exponent x A : x \in A -> #[x] %| exponent A.
+Proof. by move=> Ax; rewrite (biglcmn_sup x). Qed.
+
+Lemma expg_exponent x A : x \in A -> x ^+ exponent A = 1.
+Proof. by move=> Ax; apply/eqP; rewrite -order_dvdn dvdn_exponent. Qed.
+
+Lemma exponentS A B : A \subset B -> exponent A %| exponent B.
+Proof.
+by move=> sAB; apply/dvdn_biglcmP=> x Ax; rewrite dvdn_exponent ?(subsetP sAB).
+Qed.
+
+Lemma exponentP A n :
+ reflect (forall x, x \in A -> x ^+ n = 1) (exponent A %| n).
+Proof.
+apply: (iffP (dvdn_biglcmP _ _ _)) => eAn x Ax.
+ by apply/eqP; rewrite -order_dvdn eAn.
+by rewrite order_dvdn eAn.
+Qed.
+Implicit Arguments exponentP [A n].
+
+Lemma trivg_exponent G : (G :==: 1) = (exponent G %| 1).
+Proof.
+rewrite -subG1.
+by apply/subsetP/exponentP=> trG x /trG; rewrite expg1 => /set1P.
+Qed.
+
+Lemma exponent1 : exponent [1 gT] = 1%N.
+Proof. by apply/eqP; rewrite -dvdn1 -trivg_exponent eqxx. Qed.
+
+Lemma exponent_dvdn G : exponent G %| #|G|.
+Proof. by apply/dvdn_biglcmP=> x Gx; exact: order_dvdG. Qed.
+
+Lemma exponent_gt0 G : 0 < exponent G.
+Proof. exact: dvdn_gt0 (exponent_dvdn G). Qed.
+Hint Resolve exponent_gt0.
+
+Lemma pnat_exponent pi G : pi.-nat (exponent G) = pi.-group G.
+Proof.
+congr (_ && _); first by rewrite cardG_gt0 exponent_gt0.
+apply: eq_all_r => p; rewrite !mem_primes cardG_gt0 exponent_gt0 /=.
+apply: andb_id2l => p_pr; apply/idP/idP=> pG.
+ exact: dvdn_trans pG (exponent_dvdn G).
+by case/Cauchy: pG => // x Gx <-; exact: dvdn_exponent.
+Qed.
+
+Lemma exponentJ A x : exponent (A :^ x) = exponent A.
+Proof.
+rewrite /exponent (reindex_inj (conjg_inj x)).
+by apply: eq_big => [y | y _]; rewrite ?orderJ ?memJ_conjg.
+Qed.
+
+Lemma exponent_witness G : nilpotent G -> {x | x \in G & exponent G = #[x]}.
+Proof.
+move=> nilG; have [//=| /= x Gx max_x] := @arg_maxP _ 1 (mem G) order.
+exists x => //; apply/eqP; rewrite eqn_dvd dvdn_exponent // andbT.
+apply/dvdn_biglcmP=> y Gy; apply/dvdn_partP=> //= p.
+rewrite mem_primes => /andP[p_pr _]; have p_gt1: p > 1 := prime_gt1 p_pr.
+rewrite p_part pfactor_dvdn // -(leq_exp2l _ _ p_gt1) -!p_part.
+rewrite -(leq_pmul2r (part_gt0 p^' #[x])) partnC // -!order_constt.
+rewrite -orderM ?order_constt ?coprime_partC // ?max_x ?groupM ?groupX //.
+case/dprodP: (nilpotent_pcoreC p nilG) => _ _ cGpGp' _.
+have inGp := mem_normal_Hall (nilpotent_pcore_Hall _ nilG) (pcore_normal _ _).
+by red; rewrite -(centsP cGpGp') // inGp ?p_elt_constt ?groupX.
+Qed.
+
+Lemma exponent_cycle x : exponent <[x]> = #[x].
+Proof. by apply/eqP; rewrite eqn_dvd exponent_dvdn dvdn_exponent ?cycle_id. Qed.
+
+Lemma exponent_cyclic X : cyclic X -> exponent X = #|X|.
+Proof. by case/cyclicP=> x ->; exact: exponent_cycle. Qed.
+
+Lemma primes_exponent G : primes (exponent G) = primes (#|G|).
+Proof.
+apply/eq_primes => p; rewrite !mem_primes exponent_gt0 cardG_gt0 /=.
+by apply: andb_id2l => p_pr; apply: negb_inj; rewrite -!p'natE // pnat_exponent.
+Qed.
+
+Lemma pi_of_exponent G : \pi(exponent G) = \pi(G).
+Proof. by rewrite /pi_of primes_exponent. Qed.
+
+Lemma partn_exponentS pi H G :
+ H \subset G -> #|G|`_pi %| #|H| -> (exponent H)`_pi = (exponent G)`_pi.
+Proof.
+move=> sHG Gpi_dvd_H; apply/eqP; rewrite eqn_dvd.
+rewrite partn_dvd ?exponentS ?exponent_gt0 //=; apply/dvdn_partP=> // p.
+rewrite pi_of_part ?exponent_gt0 // => /andP[_ /= pi_p].
+have sppi: {subset (p : nat_pred) <= pi} by move=> q /eqnP->.
+have [P sylP] := Sylow_exists p H; have sPH := pHall_sub sylP.
+have{sylP} sylP: p.-Sylow(G) P.
+ rewrite pHallE (subset_trans sPH) //= (card_Hall sylP) eqn_dvd andbC.
+ by rewrite -{1}(partn_part _ sppi) !partn_dvd ?cardSg ?cardG_gt0.
+rewrite partn_part ?partn_biglcm //.
+apply: (@big_ind _ (dvdn^~ _)) => [|m n|x Gx]; first exact: dvd1n.
+ by rewrite dvdn_lcm => ->.
+rewrite -order_constt; have p_y := p_elt_constt p x; set y := x.`_p in p_y *.
+have sYG: <[y]> \subset G by rewrite cycle_subG groupX.
+have [z _ Pyz] := Sylow_Jsub sylP sYG p_y.
+rewrite (bigD1 (y ^ z)) ?(subsetP sPH) -?cycle_subG ?cycleJ //=.
+by rewrite orderJ part_pnat_id ?dvdn_lcml // (pi_pnat p_y).
+Qed.
+
+Lemma exponent_Hall pi G H : pi.-Hall(G) H -> exponent H = (exponent G)`_pi.
+Proof.
+move=> hallH; have [sHG piH _] := and3P hallH.
+rewrite -(partn_exponentS sHG) -?(card_Hall hallH) ?part_pnat_id //.
+by apply: pnat_dvd piH; exact: exponent_dvdn.
+Qed.
+
+Lemma exponent_Zgroup G : Zgroup G -> exponent G = #|G|.
+Proof.
+move/forall_inP=> ZgG; apply/eqP; rewrite eqn_dvd exponent_dvdn.
+apply/(dvdn_partP _ (cardG_gt0 _)) => p _.
+have [S sylS] := Sylow_exists p G; rewrite -(card_Hall sylS).
+have /cyclicP[x defS]: cyclic S by rewrite ZgG ?(p_Sylow sylS).
+by rewrite defS dvdn_exponent // -cycle_subG -defS (pHall_sub sylS).
+Qed.
+
+Lemma cprod_exponent A B G :
+ A \* B = G -> lcmn (exponent A) (exponent B) = (exponent G).
+Proof.
+case/cprodP=> [[K H -> ->{A B}] <- cKH].
+apply/eqP; rewrite eqn_dvd dvdn_lcm !exponentS ?mulG_subl ?mulG_subr //=.
+apply/exponentP=> _ /imset2P[x y Kx Hy ->].
+rewrite -[1]mulg1 expgMn; last by red; rewrite -(centsP cKH).
+congr (_ * _); apply/eqP; rewrite -order_dvdn.
+ by rewrite (dvdn_trans (dvdn_exponent Kx)) ?dvdn_lcml.
+by rewrite (dvdn_trans (dvdn_exponent Hy)) ?dvdn_lcmr.
+Qed.
+
+Lemma dprod_exponent A B G :
+ A \x B = G -> lcmn (exponent A) (exponent B) = (exponent G).
+Proof.
+case/dprodP=> [[K H -> ->{A B}] defG cKH _].
+by apply: cprod_exponent; rewrite cprodE.
+Qed.
+
+Lemma sub_LdivT A n : (A \subset 'Ldiv_n()) = (exponent A %| n).
+Proof. by apply/subsetP/exponentP=> eAn x /eAn; rewrite inE => /eqP. Qed.
+
+Lemma LdivT_J n x : 'Ldiv_n() :^ x = 'Ldiv_n().
+Proof.
+apply/setP=> y; rewrite !inE mem_conjg inE -conjXg.
+by rewrite (canF_eq (conjgKV x)) conj1g.
+Qed.
+
+Lemma LdivJ n A x : 'Ldiv_n(A :^ x) = 'Ldiv_n(A) :^ x.
+Proof. by rewrite conjIg LdivT_J. Qed.
+
+Lemma sub_Ldiv A n : (A \subset 'Ldiv_n(A)) = (exponent A %| n).
+Proof. by rewrite subsetI subxx sub_LdivT. Qed.
+
+Lemma group_Ldiv G n : abelian G -> group_set 'Ldiv_n(G).
+Proof.
+move=> cGG; apply/group_setP.
+split=> [|x y]; rewrite !inE ?group1 ?expg1n //=.
+case/andP=> Gx /eqP xn /andP[Gy /eqP yn].
+rewrite groupM //= expgMn ?xn ?yn ?mulg1 //; exact: (centsP cGG).
+Qed.
+
+Lemma abelian_exponent_gen A : abelian A -> exponent <<A>> = exponent A.
+Proof.
+rewrite -abelian_gen; set n := exponent A; set G := <<A>> => cGG.
+apply/eqP; rewrite eqn_dvd andbC exponentS ?subset_gen //= -sub_Ldiv.
+rewrite -(gen_set_id (group_Ldiv n cGG)) genS // subsetI subset_gen /=.
+by rewrite sub_LdivT.
+Qed.
+
+Lemma abelem_pgroup p A : p.-abelem A -> p.-group A.
+Proof. by case/andP. Qed.
+
+Lemma abelem_abelian p A : p.-abelem A -> abelian A.
+Proof. by case/and3P. Qed.
+
+Lemma abelem1 p : p.-abelem [1 gT].
+Proof. by rewrite /abelem pgroup1 abelian1 exponent1 dvd1n. Qed.
+
+Lemma abelemE p G : prime p -> p.-abelem G = abelian G && (exponent G %| p).
+Proof.
+move=> p_pr; rewrite /abelem -pnat_exponent andbA -!(andbC (_ %| _)).
+by case: (dvdn_pfactor _ 1 p_pr) => // [[k _ ->]]; rewrite pnat_exp pnat_id.
+Qed.
+
+Lemma abelemP p G :
+ prime p ->
+ reflect (abelian G /\ forall x, x \in G -> x ^+ p = 1) (p.-abelem G).
+Proof.
+by move=> p_pr; rewrite abelemE //; apply: (iffP andP) => [] [-> /exponentP].
+Qed.
+
+Lemma abelem_order_p p G x : p.-abelem G -> x \in G -> x != 1 -> #[x] = p.
+Proof.
+case/and3P=> pG _ eG Gx; rewrite -cycle_eq1 => ntX.
+have{ntX} [p_pr p_x _] := pgroup_pdiv (mem_p_elt pG Gx) ntX.
+by apply/eqP; rewrite eqn_dvd p_x andbT order_dvdn (exponentP eG).
+Qed.
+
+Lemma cyclic_abelem_prime p X : p.-abelem X -> cyclic X -> X :!=: 1 -> #|X| = p.
+Proof.
+move=> abelX cycX; case/cyclicP: cycX => x -> in abelX *.
+by rewrite cycle_eq1; exact: abelem_order_p abelX (cycle_id x).
+Qed.
+
+Lemma cycle_abelem p x : p.-elt x || prime p -> p.-abelem <[x]> = (#[x] %| p).
+Proof.
+move=> p_xVpr; rewrite /abelem cycle_abelian /=.
+apply/andP/idP=> [[_ xp1] | x_dvd_p].
+ by rewrite order_dvdn (exponentP xp1) ?cycle_id.
+split; last exact: dvdn_trans (exponent_dvdn _) x_dvd_p.
+by case/orP: p_xVpr => // /pnat_id; exact: pnat_dvd.
+Qed.
+
+Lemma exponent2_abelem G : exponent G %| 2 -> 2.-abelem G.
+Proof.
+move/exponentP=> expG; apply/abelemP=> //; split=> //.
+apply/centsP=> x Gx y Gy; apply: (mulIg x); apply: (mulgI y).
+by rewrite -!mulgA !(mulgA y) -!(expgS _ 1) !expG ?mulg1 ?groupM.
+Qed.
+
+Lemma prime_abelem p G : prime p -> #|G| = p -> p.-abelem G.
+Proof.
+move=> p_pr oG; rewrite /abelem -oG exponent_dvdn.
+by rewrite /pgroup cyclic_abelian ?prime_cyclic ?oG ?pnat_id.
+Qed.
+
+Lemma abelem_cyclic p G : p.-abelem G -> cyclic G = (logn p #|G| <= 1).
+Proof.
+move=> abelG; have [pG _ expGp] := and3P abelG.
+case: (eqsVneq G 1) => [-> | ntG]; first by rewrite cyclic1 cards1 logn1.
+have [p_pr _ [e oG]] := pgroup_pdiv pG ntG; apply/idP/idP.
+ case/cyclicP=> x defG; rewrite -(pfactorK 1 p_pr) dvdn_leq_log ?prime_gt0 //.
+ by rewrite defG order_dvdn (exponentP expGp) // defG cycle_id.
+by rewrite oG pfactorK // ltnS leqn0 => e0; rewrite prime_cyclic // oG (eqP e0).
+Qed.
+
+Lemma abelemS p H G : H \subset G -> p.-abelem G -> p.-abelem H.
+Proof.
+move=> sHG /and3P[cGG pG Gp1]; rewrite /abelem.
+by rewrite (pgroupS sHG) // (abelianS sHG) // (dvdn_trans (exponentS sHG)).
+Qed.
+
+Lemma abelemJ p G x : p.-abelem (G :^ x) = p.-abelem G.
+Proof. by rewrite /abelem pgroupJ abelianJ exponentJ. Qed.
+
+Lemma cprod_abelem p A B G :
+ A \* B = G -> p.-abelem G = p.-abelem A && p.-abelem B.
+Proof.
+case/cprodP=> [[H K -> ->{A B}] defG cHK].
+apply/idP/andP=> [abelG | []].
+ by rewrite !(abelemS _ abelG) // -defG (mulG_subl, mulG_subr).
+case/and3P=> pH cHH expHp; case/and3P=> pK cKK expKp.
+rewrite -defG /abelem pgroupM pH pK abelianM cHH cKK cHK /=.
+apply/exponentP=> _ /imset2P[x y Hx Ky ->].
+rewrite expgMn; last by red; rewrite -(centsP cHK).
+by rewrite (exponentP expHp) // (exponentP expKp) // mul1g.
+Qed.
+
+Lemma dprod_abelem p A B G :
+ A \x B = G -> p.-abelem G = p.-abelem A && p.-abelem B.
+Proof.
+move=> defG; case/dprodP: (defG) => _ _ _ tiHK.
+by apply: cprod_abelem; rewrite -dprodEcp.
+Qed.
+
+Lemma is_abelem_pgroup p G : p.-group G -> is_abelem G = p.-abelem G.
+Proof.
+rewrite /is_abelem => pG.
+case: (eqsVneq G 1) => [-> | ntG]; first by rewrite !abelem1.
+by have [p_pr _ [k ->]] := pgroup_pdiv pG ntG; rewrite pdiv_pfactor.
+Qed.
+
+Lemma is_abelemP G : reflect (exists2 p, prime p & p.-abelem G) (is_abelem G).
+Proof.
+apply: (iffP idP) => [abelG | [p p_pr abelG]].
+ case: (eqsVneq G 1) => [-> | ntG]; first by exists 2; rewrite ?abelem1.
+ by exists (pdiv #|G|); rewrite ?pdiv_prime // ltnNge -trivg_card_le1.
+by rewrite (is_abelem_pgroup (abelem_pgroup abelG)).
+Qed.
+
+Lemma pElemP p A E : reflect (E \subset A /\ p.-abelem E) (E \in 'E_p(A)).
+Proof. by rewrite inE; exact: andP. Qed.
+Implicit Arguments pElemP [p A E].
+
+Lemma pElemS p A B : A \subset B -> 'E_p(A) \subset 'E_p(B).
+Proof.
+by move=> sAB; apply/subsetP=> E; rewrite !inE => /andP[/subset_trans->].
+Qed.
+
+Lemma pElemI p A B : 'E_p(A :&: B) = 'E_p(A) :&: subgroups B.
+Proof. by apply/setP=> E; rewrite !inE subsetI andbAC. Qed.
+
+Lemma pElemJ x p A E : ((E :^ x)%G \in 'E_p(A :^ x)) = (E \in 'E_p(A)).
+Proof. by rewrite !inE conjSg abelemJ. Qed.
+
+Lemma pnElemP p n A E :
+ reflect [/\ E \subset A, p.-abelem E & logn p #|E| = n] (E \in 'E_p^n(A)).
+Proof. by rewrite !inE -andbA; apply: (iffP and3P) => [] [-> -> /eqP]. Qed.
+Implicit Arguments pnElemP [p n A E].
+
+Lemma pnElemPcard p n A E :
+ E \in 'E_p^n(A) -> [/\ E \subset A, p.-abelem E & #|E| = p ^ n]%N.
+Proof.
+by case/pnElemP=> -> abelE <-; rewrite -card_pgroup // abelem_pgroup.
+Qed.
+
+Lemma card_pnElem p n A E : E \in 'E_p^n(A) -> #|E| = (p ^ n)%N.
+Proof. by case/pnElemPcard. Qed.
+
+Lemma pnElem0 p G : 'E_p^0(G) = [set 1%G].
+Proof.
+apply/setP=> E; rewrite !inE -andbA; apply/and3P/idP=> [[_ pE] | /eqP->].
+ apply: contraLR; case/(pgroup_pdiv (abelem_pgroup pE)) => p_pr _ [k ->].
+ by rewrite pfactorK.
+by rewrite sub1G abelem1 cards1 logn1.
+Qed.
+
+Lemma pnElem_prime p n A E : E \in 'E_p^n.+1(A) -> prime p.
+Proof. by case/pnElemP=> _ _; rewrite lognE; case: prime. Qed.
+
+Lemma pnElemE p n A :
+ prime p -> 'E_p^n(A) = [set E in 'E_p(A) | #|E| == (p ^ n)%N].
+Proof.
+move/pfactorK=> pnK; apply/setP=> E; rewrite 3!inE.
+case: (@andP (E \subset A)) => //= [[_]] /andP[/p_natP[k ->] _].
+by rewrite pnK (can_eq pnK).
+Qed.
+
+Lemma pnElemS p n A B : A \subset B -> 'E_p^n(A) \subset 'E_p^n(B).
+Proof.
+move=> sAB; apply/subsetP=> E.
+by rewrite !inE -!andbA => /andP[/subset_trans->].
+Qed.
+
+Lemma pnElemI p n A B : 'E_p^n(A :&: B) = 'E_p^n(A) :&: subgroups B.
+Proof. by apply/setP=> E; rewrite !inE subsetI -!andbA; do !bool_congr. Qed.
+
+Lemma pnElemJ x p n A E : ((E :^ x)%G \in 'E_p^n(A :^ x)) = (E \in 'E_p^n(A)).
+Proof. by rewrite inE pElemJ cardJg !inE. Qed.
+
+Lemma abelem_pnElem p n G :
+ p.-abelem G -> n <= logn p #|G| -> exists E, E \in 'E_p^n(G).
+Proof.
+case: n => [|n] abelG lt_nG; first by exists 1%G; rewrite pnElem0 set11.
+have p_pr: prime p by move: lt_nG; rewrite lognE; case: prime.
+case/(normal_pgroup (abelem_pgroup abelG)): lt_nG => // E [sEG _ oE].
+by exists E; rewrite pnElemE // !inE oE sEG (abelemS sEG) /=.
+Qed.
+
+Lemma card_p1Elem p A X : X \in 'E_p^1(A) -> #|X| = p.
+Proof. exact: card_pnElem. Qed.
+
+Lemma p1ElemE p A : prime p -> 'E_p^1(A) = [set X in subgroups A | #|X| == p].
+Proof.
+move=> p_pr; apply/setP=> X; rewrite pnElemE // !inE -andbA; congr (_ && _).
+by apply: andb_idl => /eqP oX; rewrite prime_abelem ?oX.
+Qed.
+
+Lemma TIp1ElemP p A X Y :
+ X \in 'E_p^1(A) -> Y \in 'E_p^1(A) -> reflect (X :&: Y = 1) (X :!=: Y).
+Proof.
+move=> EpX EpY; have p_pr := pnElem_prime EpX.
+have [oX oY] := (card_p1Elem EpX, card_p1Elem EpY).
+have [<- |] := altP eqP.
+ by right=> X1; rewrite -oX -(setIid X) X1 cards1 in p_pr.
+by rewrite eqEcard oX oY leqnn andbT; left; rewrite prime_TIg ?oX.
+Qed.
+
+Lemma card_p1Elem_pnElem p n A E :
+ E \in 'E_p^n(A) -> #|'E_p^1(E)| = (\sum_(i < n) p ^ i)%N.
+Proof.
+case/pnElemP=> _ {A} abelE dimE; have [pE cEE _] := and3P abelE.
+have [E1 | ntE] := eqsVneq E 1.
+ rewrite -dimE E1 cards1 logn1 big_ord0 eq_card0 // => X.
+ by rewrite !inE subG1 trivg_card1; case: eqP => // ->; rewrite logn1 andbF.
+have [p_pr _ _] := pgroup_pdiv pE ntE; have p_gt1 := prime_gt1 p_pr.
+apply/eqP; rewrite -(@eqn_pmul2l (p - 1)) ?subn_gt0 // subn1 -predn_exp.
+have groupD1_inj: injective (fun X => (gval X)^#).
+ apply: can_inj (@generated_group _) _ => X.
+ by apply: val_inj; rewrite /= genD1 ?group1 ?genGid.
+rewrite -dimE -card_pgroup // (cardsD1 1 E) group1 /= mulnC.
+rewrite -(card_imset _ groupD1_inj) eq_sym.
+apply/eqP; apply: card_uniform_partition => [X'|].
+ case/imsetP=> X; rewrite pnElemE // expn1 => /setIdP[_ /eqP <-] ->.
+ by rewrite (cardsD1 1 X) group1.
+apply/and3P; split; last 1 first.
+- apply/imsetP=> [[X /card_p1Elem oX X'0]].
+ by rewrite -oX (cardsD1 1) -X'0 group1 cards0 in p_pr.
+- rewrite eqEsubset; apply/andP; split.
+ by apply/bigcupsP=> _ /imsetP[X /pnElemP[sXE _ _] ->]; exact: setSD.
+ apply/subsetP=> x /setD1P[ntx Ex].
+ apply/bigcupP; exists <[x]>^#; last by rewrite !inE ntx cycle_id.
+ apply/imsetP; exists <[x]>%G; rewrite ?p1ElemE // !inE cycle_subG Ex /=.
+ by rewrite -orderE (abelem_order_p abelE).
+apply/trivIsetP=> _ _ /imsetP[X EpX ->] /imsetP[Y EpY ->]; apply/implyP.
+rewrite (inj_eq groupD1_inj) -setI_eq0 -setDIl setD_eq0 subG1.
+by rewrite (sameP eqP (TIp1ElemP EpX EpY)) implybb.
+Qed.
+
+Lemma card_p1Elem_p2Elem p A E : E \in 'E_p^2(A) -> #|'E_p^1(E)| = p.+1.
+Proof. by move/card_p1Elem_pnElem->; rewrite big_ord_recl big_ord1. Qed.
+
+Lemma p2Elem_dprodP p A E X Y :
+ E \in 'E_p^2(A) -> X \in 'E_p^1(E) -> Y \in 'E_p^1(E) ->
+ reflect (X \x Y = E) (X :!=: Y).
+Proof.
+move=> Ep2E EpX EpY; have [_ abelE oE] := pnElemPcard Ep2E.
+apply: (iffP (TIp1ElemP EpX EpY)) => [tiXY|]; last by case/dprodP.
+have [[sXE _ oX] [sYE _ oY]] := (pnElemPcard EpX, pnElemPcard EpY).
+rewrite dprodE ?(sub_abelian_cent2 (abelem_abelian abelE)) //.
+by apply/eqP; rewrite eqEcard mul_subG //= TI_cardMg // oX oY oE.
+Qed.
+
+Lemma nElemP n G E : reflect (exists p, E \in 'E_p^n(G)) (E \in 'E^n(G)).
+Proof.
+rewrite ['E^n(G)]big_mkord.
+apply: (iffP bigcupP) => [[[p /= _] _] | [p]]; first by exists p.
+case: n => [|n EpnE]; first by rewrite pnElem0; exists ord0; rewrite ?pnElem0.
+suffices lepG: p < #|G|.+1 by exists (Ordinal lepG).
+have:= EpnE; rewrite pnElemE ?(pnElem_prime EpnE) // !inE -andbA ltnS.
+case/and3P=> sEG _ oE; rewrite dvdn_leq // (dvdn_trans _ (cardSg sEG)) //.
+by rewrite (eqP oE) dvdn_exp.
+Qed.
+Implicit Arguments nElemP [n G E].
+
+Lemma nElem0 G : 'E^0(G) = [set 1%G].
+Proof.
+apply/setP=> E; apply/nElemP/idP=> [[p] |]; first by rewrite pnElem0.
+by exists 2; rewrite pnElem0.
+Qed.
+
+Lemma nElem1P G E :
+ reflect (E \subset G /\ exists2 p, prime p & #|E| = p) (E \in 'E^1(G)).
+Proof.
+apply: (iffP nElemP) => [[p pE] | [sEG [p p_pr oE]]].
+ have p_pr := pnElem_prime pE; rewrite pnElemE // !inE -andbA in pE.
+ by case/and3P: pE => -> _ /eqP; split; last exists p.
+exists p; rewrite pnElemE // !inE sEG oE eqxx abelemE // -oE exponent_dvdn.
+by rewrite cyclic_abelian // prime_cyclic // oE.
+Qed.
+
+Lemma nElemS n G H : G \subset H -> 'E^n(G) \subset 'E^n(H).
+Proof.
+move=> sGH; apply/subsetP=> E /nElemP[p EpnG_E].
+by apply/nElemP; exists p; rewrite // (subsetP (pnElemS _ _ sGH)).
+Qed.
+
+Lemma nElemI n G H : 'E^n(G :&: H) = 'E^n(G) :&: subgroups H.
+Proof.
+apply/setP=> E; apply/nElemP/setIP=> [[p] | []].
+ by rewrite pnElemI; case/setIP; split=> //; apply/nElemP; exists p.
+by case/nElemP=> p EpnG_E sHE; exists p; rewrite pnElemI inE EpnG_E.
+Qed.
+
+Lemma def_pnElem p n G : 'E_p^n(G) = 'E_p(G) :&: 'E^n(G).
+Proof.
+apply/setP=> E; rewrite inE in_setI; apply: andb_id2l => /pElemP[sEG abelE].
+apply/idP/nElemP=> [|[q]]; first by exists p; rewrite !inE sEG abelE.
+rewrite !inE -2!andbA => /and4P[_ /pgroupP qE _].
+case: (eqVneq E 1%G) => [-> | ]; first by rewrite cards1 !logn1.
+case/(pgroup_pdiv (abelem_pgroup abelE)) => p_pr pE _.
+by rewrite (eqnP (qE p p_pr pE)).
+Qed.
+
+Lemma pmaxElemP p A E :
+ reflect (E \in 'E_p(A) /\ forall H, H \in 'E_p(A) -> E \subset H -> H :=: E)
+ (E \in 'E*_p(A)).
+Proof. by rewrite [E \in 'E*_p(A)]inE; exact: (iffP maxgroupP). Qed.
+
+Lemma pmaxElem_exists p A D :
+ D \in 'E_p(A) -> {E | E \in 'E*_p(A) & D \subset E}.
+Proof.
+move=> EpD; have [E maxE sDE] := maxgroup_exists (EpD : mem 'E_p(A) D).
+by exists E; rewrite // inE.
+Qed.
+
+Lemma pmaxElem_LdivP p G E :
+ prime p -> reflect ('Ldiv_p('C_G(E)) = E) (E \in 'E*_p(G)).
+Proof.
+move=> p_pr; apply: (iffP (pmaxElemP p G E)) => [[] | defE].
+ case/pElemP=> sEG abelE maxE; have [_ cEE eE] := and3P abelE.
+ apply/setP=> x; rewrite !inE -andbA; apply/and3P/idP=> [[Gx cEx xp] | Ex].
+ rewrite -(maxE (<[x]> <*> E)%G) ?joing_subr //.
+ by rewrite -cycle_subG joing_subl.
+ rewrite inE join_subG cycle_subG Gx sEG /=.
+ rewrite (cprod_abelem _ (cprodEY _)); last by rewrite centsC cycle_subG.
+ by rewrite cycle_abelem ?p_pr ?orbT // order_dvdn xp.
+ by rewrite (subsetP sEG) // (subsetP cEE) // (exponentP eE).
+split=> [|H]; last first.
+ case/pElemP=> sHG /abelemP[// | cHH Hp1] sEH.
+ apply/eqP; rewrite eqEsubset sEH andbC /= -defE; apply/subsetP=> x Hx.
+ by rewrite 3!inE (subsetP sHG) // Hp1 ?(subsetP (centsS _ cHH)) /=.
+apply/pElemP; split; first by rewrite -defE -setIA subsetIl.
+apply/abelemP=> //; rewrite /abelian -{1 3}defE setIAC subsetIr.
+by split=> //; apply/exponentP; rewrite -sub_LdivT setIAC subsetIr.
+Qed.
+
+Lemma pmaxElemS p A B :
+ A \subset B -> 'E*_p(B) :&: subgroups A \subset 'E*_p(A).
+Proof.
+move=> sAB; apply/subsetP=> E; rewrite !inE.
+case/andP=> /maxgroupP[/pElemP[_ abelE] maxE] sEA.
+apply/maxgroupP; rewrite inE sEA; split=> // D EpD.
+by apply: maxE; apply: subsetP EpD; exact: pElemS.
+Qed.
+
+Lemma pmaxElemJ p A E x : ((E :^ x)%G \in 'E*_p(A :^ x)) = (E \in 'E*_p(A)).
+Proof.
+apply/pmaxElemP/pmaxElemP=> [] [EpE maxE].
+ rewrite pElemJ in EpE; split=> //= H EpH sEH; apply: (act_inj 'Js x).
+ by apply: maxE; rewrite ?conjSg ?pElemJ.
+rewrite pElemJ; split=> // H; rewrite -(actKV 'JG x H) pElemJ conjSg => EpHx'.
+by move/maxE=> /= ->.
+Qed.
+
+Lemma grank_min B : 'm(<<B>>) <= #|B|.
+Proof.
+by rewrite /gen_rank; case: arg_minP => [|_ _ -> //]; rewrite genGid.
+Qed.
+
+Lemma grank_witness G : {B | <<B>> = G & #|B| = 'm(G)}.
+Proof.
+rewrite /gen_rank; case: arg_minP => [|B defG _]; first by rewrite genGid.
+by exists B; first exact/eqP.
+Qed.
+
+Lemma p_rank_witness p G : {E | E \in 'E_p^('r_p(G))(G)}.
+Proof.
+have [E EG_E mE]: {E | E \in 'E_p(G) & 'r_p(G) = logn p #|E| }.
+ by apply: eq_bigmax_cond; rewrite (cardD1 1%G) inE sub1G abelem1.
+by exists E; rewrite inE EG_E -mE /=.
+Qed.
+
+Lemma p_rank_geP p n G : reflect (exists E, E \in 'E_p^n(G)) (n <= 'r_p(G)).
+Proof.
+apply: (iffP idP) => [|[E]]; last first.
+ by rewrite inE => /andP[Ep_E /eqP <-]; rewrite (bigmax_sup E).
+have [D /pnElemP[sDG abelD <-]] := p_rank_witness p G.
+by case/abelem_pnElem=> // E; exists E; exact: (subsetP (pnElemS _ _ sDG)).
+Qed.
+
+Lemma p_rank_gt0 p H : ('r_p(H) > 0) = (p \in \pi(H)).
+Proof.
+rewrite mem_primes cardG_gt0 /=; apply/p_rank_geP/andP=> [[E] | [p_pr]].
+ case/pnElemP=> sEG _; rewrite lognE; case: and3P => // [[-> _ pE] _].
+ by rewrite (dvdn_trans _ (cardSg sEG)).
+case/Cauchy=> // x Hx ox; exists <[x]>%G; rewrite 2!inE [#|_|]ox cycle_subG.
+by rewrite Hx (pfactorK 1) ?abelemE // cycle_abelian -ox exponent_dvdn.
+Qed.
+
+Lemma p_rank1 p : 'r_p([1 gT]) = 0.
+Proof. by apply/eqP; rewrite eqn0Ngt p_rank_gt0 /= cards1. Qed.
+
+Lemma logn_le_p_rank p A E : E \in 'E_p(A) -> logn p #|E| <= 'r_p(A).
+Proof. by move=> EpA_E; rewrite (bigmax_sup E). Qed.
+
+Lemma p_rank_le_logn p G : 'r_p(G) <= logn p #|G|.
+Proof.
+have [E EpE] := p_rank_witness p G.
+by have [sEG _ <-] := pnElemP EpE; exact: lognSg.
+Qed.
+
+Lemma p_rank_abelem p G : p.-abelem G -> 'r_p(G) = logn p #|G|.
+Proof.
+move=> abelG; apply/eqP; rewrite eqn_leq andbC (bigmax_sup G) //.
+ by apply/bigmax_leqP=> E; rewrite inE => /andP[/lognSg->].
+by rewrite inE subxx.
+Qed.
+
+Lemma p_rankS p A B : A \subset B -> 'r_p(A) <= 'r_p(B).
+Proof.
+move=> sAB; apply/bigmax_leqP=> E /(subsetP (pElemS p sAB)) EpB_E.
+by rewrite (bigmax_sup E).
+Qed.
+
+Lemma p_rankElem_max p A : 'E_p^('r_p(A))(A) \subset 'E*_p(A).
+Proof.
+apply/subsetP=> E /setIdP[EpE dimE].
+apply/pmaxElemP; split=> // F EpF sEF; apply/eqP.
+have pF: p.-group F by case/pElemP: EpF => _ /and3P[].
+have pE: p.-group E by case/pElemP: EpE => _ /and3P[].
+rewrite eq_sym eqEcard sEF dvdn_leq // (card_pgroup pE) (card_pgroup pF).
+by rewrite (eqP dimE) dvdn_exp2l // logn_le_p_rank.
+Qed.
+
+Lemma p_rankJ p A x : 'r_p(A :^ x) = 'r_p(A).
+Proof.
+rewrite /p_rank (reindex_inj (act_inj 'JG x)).
+by apply: eq_big => [E | E _]; rewrite ?cardJg ?pElemJ.
+Qed.
+
+Lemma p_rank_Sylow p G H : p.-Sylow(G) H -> 'r_p(H) = 'r_p(G).
+Proof.
+move=> sylH; apply/eqP; rewrite eqn_leq (p_rankS _ (pHall_sub sylH)) /=.
+apply/bigmax_leqP=> E; rewrite inE => /andP[sEG abelE].
+have [P sylP sEP] := Sylow_superset sEG (abelem_pgroup abelE).
+have [x _ ->] := Sylow_trans sylP sylH.
+by rewrite p_rankJ -(p_rank_abelem abelE) (p_rankS _ sEP).
+Qed.
+
+Lemma p_rank_Hall pi p G H : pi.-Hall(G) H -> p \in pi -> 'r_p(H) = 'r_p(G).
+Proof.
+move=> hallH pi_p; have [P sylP] := Sylow_exists p H.
+by rewrite -(p_rank_Sylow sylP) (p_rank_Sylow (subHall_Sylow hallH pi_p sylP)).
+Qed.
+
+Lemma p_rank_pmaxElem_exists p r G :
+ 'r_p(G) >= r -> exists2 E, E \in 'E*_p(G) & 'r_p(E) >= r.
+Proof.
+case/p_rank_geP=> D /setIdP[EpD /eqP <- {r}].
+have [E EpE sDE] := pmaxElem_exists EpD; exists E => //.
+case/pmaxElemP: EpE => /setIdP[_ abelE] _.
+by rewrite (p_rank_abelem abelE) lognSg.
+Qed.
+
+Lemma rank1 : 'r([1 gT]) = 0.
+Proof. by rewrite ['r(1)]big1_seq // => p _; rewrite p_rank1. Qed.
+
+Lemma p_rank_le_rank p G : 'r_p(G) <= 'r(G).
+Proof.
+case: (posnP 'r_p(G)) => [-> //|]; rewrite p_rank_gt0 mem_primes.
+case/and3P=> p_pr _ pG; have lepg: p < #|G|.+1 by rewrite ltnS dvdn_leq.
+by rewrite ['r(G)]big_mkord (bigmax_sup (Ordinal lepg)).
+Qed.
+
+Lemma rank_gt0 G : ('r(G) > 0) = (G :!=: 1).
+Proof.
+case: (eqsVneq G 1) => [-> |]; first by rewrite rank1 eqxx.
+case: (trivgVpdiv G) => [-> | [p p_pr]]; first by case/eqP.
+case/Cauchy=> // x Gx oxp ->; apply: leq_trans (p_rank_le_rank p G).
+have EpGx: <[x]>%G \in 'E_p(G).
+ by rewrite inE cycle_subG Gx abelemE // cycle_abelian -oxp exponent_dvdn.
+by apply: leq_trans (logn_le_p_rank EpGx); rewrite -orderE oxp logn_prime ?eqxx.
+Qed.
+
+Lemma rank_witness G : {p | prime p & 'r(G) = 'r_p(G)}.
+Proof.
+have [p _ defmG]: {p : 'I_(#|G|.+1) | true & 'r(G) = 'r_p(G)}.
+ by rewrite ['r(G)]big_mkord; apply: eq_bigmax_cond; rewrite card_ord.
+case: (eqsVneq G 1) => [-> | ]; first by exists 2; rewrite // rank1 p_rank1.
+by rewrite -rank_gt0 defmG p_rank_gt0 mem_primes; case/andP; exists p.
+Qed.
+
+Lemma rank_pgroup p G : p.-group G -> 'r(G) = 'r_p(G).
+Proof.
+move=> pG; apply/eqP; rewrite eqn_leq p_rank_le_rank andbT.
+rewrite ['r(G)]big_mkord; apply/bigmax_leqP=> [[q /= _] _].
+case: (posnP 'r_q(G)) => [-> // |]; rewrite p_rank_gt0 mem_primes.
+by case/and3P=> q_pr _ qG; rewrite (eqnP (pgroupP pG q q_pr qG)).
+Qed.
+
+Lemma rank_Sylow p G P : p.-Sylow(G) P -> 'r(P) = 'r_p(G).
+Proof.
+move=> sylP; have pP := pHall_pgroup sylP.
+by rewrite -(p_rank_Sylow sylP) -(rank_pgroup pP).
+Qed.
+
+Lemma rank_abelem p G : p.-abelem G -> 'r(G) = logn p #|G|.
+Proof.
+by move=> abelG; rewrite (rank_pgroup (abelem_pgroup abelG)) p_rank_abelem.
+Qed.
+
+Lemma nt_pnElem p n E A : E \in 'E_p^n(A) -> n > 0 -> E :!=: 1.
+Proof. by case/pnElemP=> _ /rank_abelem <- <-; rewrite rank_gt0. Qed.
+
+Lemma rankJ A x : 'r(A :^ x) = 'r(A).
+Proof. by rewrite /rank cardJg; apply: eq_bigr => p _; rewrite p_rankJ. Qed.
+
+Lemma rankS A B : A \subset B -> 'r(A) <= 'r(B).
+Proof.
+move=> sAB; rewrite /rank !big_mkord; apply/bigmax_leqP=> p _.
+have leAB: #|A| < #|B|.+1 by rewrite ltnS subset_leq_card.
+by rewrite (bigmax_sup (widen_ord leAB p)) // p_rankS.
+Qed.
+
+Lemma rank_geP n G : reflect (exists E, E \in 'E^n(G)) (n <= 'r(G)).
+Proof.
+apply: (iffP idP) => [|[E]].
+ have [p _ ->] := rank_witness G; case/p_rank_geP=> E.
+ by rewrite def_pnElem; case/setIP; exists E.
+case/nElemP=> p; rewrite inE => /andP[EpG_E /eqP <-].
+by rewrite (leq_trans (logn_le_p_rank EpG_E)) ?p_rank_le_rank.
+Qed.
+
+End ExponentAbelem.
+
+Implicit Arguments LdivP [gT A n x].
+Implicit Arguments exponentP [gT A n].
+Implicit Arguments abelemP [gT p G].
+Implicit Arguments is_abelemP [gT G].
+Implicit Arguments pElemP [gT p A E].
+Implicit Arguments pnElemP [gT p n A E].
+Implicit Arguments nElemP [gT n G E].
+Implicit Arguments nElem1P [gT G E].
+Implicit Arguments pmaxElemP [gT p A E].
+Implicit Arguments pmaxElem_LdivP [gT p G E].
+Implicit Arguments p_rank_geP [gT p n G].
+Implicit Arguments rank_geP [gT n G].
+
+Section MorphAbelem.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Implicit Types (G H E : {group aT}) (A B : {set aT}).
+
+Lemma exponent_morphim G : exponent (f @* G) %| exponent G.
+Proof.
+apply/exponentP=> _ /morphimP[x Dx Gx ->].
+by rewrite -morphX // expg_exponent // morph1.
+Qed.
+
+Lemma morphim_LdivT n : f @* 'Ldiv_n() \subset 'Ldiv_n().
+Proof.
+apply/subsetP=> _ /morphimP[x Dx xn ->]; rewrite inE in xn.
+by rewrite inE -morphX // (eqP xn) morph1.
+Qed.
+
+Lemma morphim_Ldiv n A : f @* 'Ldiv_n(A) \subset 'Ldiv_n(f @* A).
+Proof.
+by apply: subset_trans (morphimI f A _) (setIS _ _); exact: morphim_LdivT.
+Qed.
+
+Lemma morphim_abelem p G : p.-abelem G -> p.-abelem (f @* G).
+Proof.
+case: (eqsVneq G 1) => [-> | ntG] abelG; first by rewrite morphim1 abelem1.
+have [p_pr _ _] := pgroup_pdiv (abelem_pgroup abelG) ntG.
+case/abelemP: abelG => // abG elemG; apply/abelemP; rewrite ?morphim_abelian //.
+by split=> // _ /morphimP[x Dx Gx ->]; rewrite -morphX // elemG ?morph1.
+Qed.
+
+Lemma morphim_pElem p G E : E \in 'E_p(G) -> (f @* E)%G \in 'E_p(f @* G).
+Proof.
+by rewrite !inE => /andP[sEG abelE]; rewrite morphimS // morphim_abelem.
+Qed.
+
+Lemma morphim_pnElem p n G E :
+ E \in 'E_p^n(G) -> {m | m <= n & (f @* E)%G \in 'E_p^m(f @* G)}.
+Proof.
+rewrite inE => /andP[EpE /eqP <-].
+by exists (logn p #|f @* E|); rewrite ?logn_morphim // inE morphim_pElem /=.
+Qed.
+
+Lemma morphim_grank G : G \subset D -> 'm(f @* G) <= 'm(G).
+Proof.
+have [B defG <-] := grank_witness G; rewrite -defG gen_subG => sBD.
+by rewrite morphim_gen ?morphimEsub ?(leq_trans (grank_min _)) ?leq_imset_card.
+Qed.
+
+(* There are no general morphism relations for the p-rank. We later prove *)
+(* some relations for the p-rank of a quotient in the QuotientAbelem section. *)
+
+End MorphAbelem.
+
+Section InjmAbelem.
+
+Variables (aT rT : finGroupType) (D G : {group aT}) (f : {morphism D >-> rT}).
+Hypotheses (injf : 'injm f) (sGD : G \subset D).
+Let defG : invm injf @* (f @* G) = G := morphim_invm injf sGD.
+
+Lemma exponent_injm : exponent (f @* G) = exponent G.
+Proof. by apply/eqP; rewrite eqn_dvd -{3}defG !exponent_morphim. Qed.
+
+Lemma injm_Ldiv n A : f @* 'Ldiv_n(A) = 'Ldiv_n(f @* A).
+Proof.
+apply/eqP; rewrite eqEsubset morphim_Ldiv.
+rewrite -[f @* 'Ldiv_n(A)](morphpre_invm injf).
+rewrite -sub_morphim_pre; last by rewrite subIset ?morphim_sub.
+rewrite injmI ?injm_invm // setISS ?morphim_LdivT //.
+by rewrite sub_morphim_pre ?morphim_sub // morphpre_invm.
+Qed.
+
+Lemma injm_abelem p : p.-abelem (f @* G) = p.-abelem G.
+Proof. by apply/idP/idP; first rewrite -{2}defG; exact: morphim_abelem. Qed.
+
+Lemma injm_pElem p (E : {group aT}) :
+ E \subset D -> ((f @* E)%G \in 'E_p(f @* G)) = (E \in 'E_p(G)).
+Proof.
+move=> sED; apply/idP/idP=> EpE; last exact: morphim_pElem.
+by rewrite -defG -(group_inj (morphim_invm injf sED)) morphim_pElem.
+Qed.
+
+Lemma injm_pnElem p n (E : {group aT}) :
+ E \subset D -> ((f @* E)%G \in 'E_p^n(f @* G)) = (E \in 'E_p^n(G)).
+Proof. by move=> sED; rewrite inE injm_pElem // card_injm ?inE. Qed.
+
+Lemma injm_nElem n (E : {group aT}) :
+ E \subset D -> ((f @* E)%G \in 'E^n(f @* G)) = (E \in 'E^n(G)).
+Proof.
+move=> sED; apply/nElemP/nElemP=> [] [p EpE];
+ by exists p; rewrite injm_pnElem in EpE *.
+Qed.
+
+Lemma injm_pmaxElem p (E : {group aT}) :
+ E \subset D -> ((f @* E)%G \in 'E*_p(f @* G)) = (E \in 'E*_p(G)).
+Proof.
+move=> sED; have defE := morphim_invm injf sED.
+apply/pmaxElemP/pmaxElemP=> [] [EpE maxE].
+ split=> [|H EpH sEH]; first by rewrite injm_pElem in EpE.
+ have sHD: H \subset D by apply: subset_trans (sGD); case/pElemP: EpH.
+ by rewrite -(morphim_invm injf sHD) [f @* H]maxE ?morphimS ?injm_pElem.
+rewrite injm_pElem //; split=> // fH Ep_fH sfEH; have [sfHG _] := pElemP Ep_fH.
+have sfHD : fH \subset f @* D by rewrite (subset_trans sfHG) ?morphimS.
+rewrite -(morphpreK sfHD); congr (f @* _).
+rewrite [_ @*^-1 fH]maxE -?sub_morphim_pre //.
+by rewrite -injm_pElem ?subsetIl // (group_inj (morphpreK sfHD)).
+Qed.
+
+Lemma injm_grank : 'm(f @* G) = 'm(G).
+Proof. by apply/eqP; rewrite eqn_leq -{3}defG !morphim_grank ?morphimS. Qed.
+
+Lemma injm_p_rank p : 'r_p(f @* G) = 'r_p(G).
+Proof.
+apply/eqP; rewrite eqn_leq; apply/andP; split.
+ have [fE] := p_rank_witness p (f @* G); move: 'r_p(_) => n Ep_fE.
+ apply/p_rank_geP; exists (f @*^-1 fE)%G.
+ rewrite -injm_pnElem ?subsetIl ?(group_inj (morphpreK _)) //.
+ by case/pnElemP: Ep_fE => sfEG _ _; rewrite (subset_trans sfEG) ?morphimS.
+have [E] := p_rank_witness p G; move: 'r_p(_) => n EpE.
+apply/p_rank_geP; exists (f @* E)%G; rewrite injm_pnElem //.
+by case/pnElemP: EpE => sEG _ _; rewrite (subset_trans sEG).
+Qed.
+
+Lemma injm_rank : 'r(f @* G) = 'r(G).
+Proof.
+apply/eqP; rewrite eqn_leq; apply/andP; split.
+ by have [p _ ->] := rank_witness (f @* G); rewrite injm_p_rank p_rank_le_rank.
+by have [p _ ->] := rank_witness G; rewrite -injm_p_rank p_rank_le_rank.
+Qed.
+
+End InjmAbelem.
+
+Section IsogAbelem.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (H : {group rT}).
+Hypothesis isoGH : G \isog H.
+
+Lemma exponent_isog : exponent G = exponent H.
+Proof. by case/isogP: isoGH => f injf <-; rewrite exponent_injm. Qed.
+
+Lemma isog_abelem p : p.-abelem G = p.-abelem H.
+Proof. by case/isogP: isoGH => f injf <-; rewrite injm_abelem. Qed.
+
+Lemma isog_grank : 'm(G) = 'm(H).
+Proof. by case/isogP: isoGH => f injf <-; rewrite injm_grank. Qed.
+
+Lemma isog_p_rank p : 'r_p(G) = 'r_p(H).
+Proof. by case/isogP: isoGH => f injf <-; rewrite injm_p_rank. Qed.
+
+Lemma isog_rank : 'r(G) = 'r(H).
+Proof. by case/isogP: isoGH => f injf <-; rewrite injm_rank. Qed.
+
+End IsogAbelem.
+
+Section QuotientAbelem.
+
+Variables (gT : finGroupType) (p : nat).
+Implicit Types E G K H : {group gT}.
+
+Lemma exponent_quotient G H : exponent (G / H) %| exponent G.
+Proof. exact: exponent_morphim. Qed.
+
+Lemma quotient_LdivT n H : 'Ldiv_n() / H \subset 'Ldiv_n().
+Proof. exact: morphim_LdivT. Qed.
+
+Lemma quotient_Ldiv n A H : 'Ldiv_n(A) / H \subset 'Ldiv_n(A / H).
+Proof. exact: morphim_Ldiv. Qed.
+
+Lemma quotient_abelem G H : p.-abelem G -> p.-abelem (G / H).
+Proof. exact: morphim_abelem. Qed.
+
+Lemma quotient_pElem G H E : E \in 'E_p(G) -> (E / H)%G \in 'E_p(G / H).
+Proof. exact: morphim_pElem. Qed.
+
+Lemma logn_quotient G H : logn p #|G / H| <= logn p #|G|.
+Proof. exact: logn_morphim. Qed.
+
+Lemma quotient_pnElem G H n E :
+ E \in 'E_p^n(G) -> {m | m <= n & (E / H)%G \in 'E_p^m(G / H)}.
+Proof. exact: morphim_pnElem. Qed.
+
+Lemma quotient_grank G H : G \subset 'N(H) -> 'm(G / H) <= 'm(G).
+Proof. exact: morphim_grank. Qed.
+
+Lemma p_rank_quotient G H : G \subset 'N(H) -> 'r_p(G) - 'r_p(H) <= 'r_p(G / H).
+Proof.
+move=> nHG; rewrite leq_subLR.
+have [E EpE] := p_rank_witness p G; have{EpE} [sEG abelE <-] := pnElemP EpE.
+rewrite -(LagrangeI E H) lognM ?cardG_gt0 //.
+rewrite -card_quotient ?(subset_trans sEG) // leq_add ?logn_le_p_rank // !inE.
+ by rewrite subsetIr (abelemS (subsetIl E H)).
+by rewrite quotientS ?quotient_abelem.
+Qed.
+
+Lemma p_rank_dprod K H G : K \x H = G -> 'r_p(K) + 'r_p(H) = 'r_p(G).
+Proof.
+move=> defG; apply/eqP; rewrite eqn_leq -leq_subLR andbC.
+have [_ defKH cKH tiKH] := dprodP defG; have nKH := cents_norm cKH.
+rewrite {1}(isog_p_rank (quotient_isog nKH tiKH)) /= -quotientMidl defKH.
+rewrite p_rank_quotient; last by rewrite -defKH mul_subG ?normG.
+have [[E EpE] [F EpF]] := (p_rank_witness p K, p_rank_witness p H).
+have [[sEK abelE <-] [sFH abelF <-]] := (pnElemP EpE, pnElemP EpF).
+have defEF: E \x F = E <*> F.
+ by rewrite dprodEY ?(centSS sFH sEK) //; apply/trivgP; rewrite -tiKH setISS.
+apply/p_rank_geP; exists (E <*> F)%G; rewrite !inE (dprod_abelem p defEF).
+rewrite -lognM ?cargG_gt0 // (dprod_card defEF) abelE abelF eqxx.
+by rewrite -(genGid G) -defKH genM_join genS ?setUSS.
+Qed.
+
+Lemma p_rank_p'quotient G H :
+ (p : nat)^'.-group H -> G \subset 'N(H) -> 'r_p(G / H) = 'r_p(G).
+Proof.
+move=> p'H nHG; have [P sylP] := Sylow_exists p G.
+have [sPG pP _] := and3P sylP; have nHP := subset_trans sPG nHG.
+have tiHP: H :&: P = 1 := coprime_TIg (p'nat_coprime p'H pP).
+rewrite -(p_rank_Sylow sylP) -(p_rank_Sylow (quotient_pHall nHP sylP)).
+by rewrite (isog_p_rank (quotient_isog nHP tiHP)).
+Qed.
+
+End QuotientAbelem.
+
+Section OhmProps.
+
+Section Generic.
+
+Variables (n : nat) (gT : finGroupType).
+Implicit Types (p : nat) (x : gT) (rT : finGroupType).
+Implicit Types (A B : {set gT}) (D G H : {group gT}).
+
+Lemma Ohm_sub G : 'Ohm_n(G) \subset G.
+Proof. by rewrite gen_subG; apply/subsetP=> x /setIdP[]. Qed.
+
+Lemma Ohm1 : 'Ohm_n([1 gT]) = 1. Proof. exact: (trivgP (Ohm_sub _)). Qed.
+
+Lemma Ohm_id G : 'Ohm_n('Ohm_n(G)) = 'Ohm_n(G).
+Proof.
+apply/eqP; rewrite eqEsubset Ohm_sub genS //.
+by apply/subsetP=> x /setIdP[Gx oxn]; rewrite inE mem_gen // inE Gx.
+Qed.
+
+Lemma Ohm_cont rT G (f : {morphism G >-> rT}) :
+ f @* 'Ohm_n(G) \subset 'Ohm_n(f @* G).
+Proof.
+rewrite morphim_gen ?genS //; last by rewrite -gen_subG Ohm_sub.
+apply/subsetP=> fx /morphimP[x Gx]; rewrite inE Gx /=.
+case/OhmPredP=> p p_pr xpn_1 -> {fx}.
+rewrite inE morphimEdom mem_imset //=; apply/OhmPredP; exists p => //.
+by rewrite -morphX // xpn_1 morph1.
+Qed.
+
+Lemma OhmS H G : H \subset G -> 'Ohm_n(H) \subset 'Ohm_n(G).
+Proof.
+move=> sHG; apply: genS; apply/subsetP=> x; rewrite !inE => /andP[Hx ->].
+by rewrite (subsetP sHG).
+Qed.
+
+Lemma OhmE p G : p.-group G -> 'Ohm_n(G) = <<'Ldiv_(p ^ n)(G)>>.
+Proof.
+move=> pG; congr <<_>>; apply/setP=> x; rewrite !inE; apply: andb_id2l => Gx.
+case: (eqVneq x 1) => [-> | ntx]; first by rewrite !expg1n.
+by rewrite (pdiv_p_elt (mem_p_elt pG Gx)).
+Qed.
+
+Lemma OhmEabelian p G :
+ p.-group G -> abelian 'Ohm_n(G) -> 'Ohm_n(G) = 'Ldiv_(p ^ n)(G).
+Proof.
+move=> pG; rewrite (OhmE pG) abelian_gen => cGGn; rewrite gen_set_id //.
+rewrite -(setIidPr (subset_gen 'Ldiv_(p ^ n)(G))) setIA.
+by rewrite [_ :&: G](setIidPl _) ?gen_subG ?subsetIl // group_Ldiv ?abelian_gen.
+Qed.
+
+Lemma Ohm_p_cycle p x :
+ p.-elt x -> 'Ohm_n(<[x]>) = <[x ^+ (p ^ (logn p #[x] - n))]>.
+Proof.
+move=> p_x; apply/eqP; rewrite (OhmE p_x) eqEsubset cycle_subG mem_gen.
+ rewrite gen_subG andbT; apply/subsetP=> y /LdivP[x_y ypn].
+ case: (leqP (logn p #[x]) n) => [|lt_n_x].
+ by rewrite -subn_eq0 => /eqP->.
+ have p_pr: prime p by move: lt_n_x; rewrite lognE; case: (prime p).
+ have def_y: <[y]> = <[x ^+ (#[x] %/ #[y])]>.
+ apply: congr_group; apply/set1P.
+ by rewrite -cycle_sub_group ?cardSg ?inE ?cycle_subG ?x_y /=.
+ rewrite -cycle_subG def_y cycle_subG -{1}(part_pnat_id p_x) p_part.
+ rewrite -{1}(subnK (ltnW lt_n_x)) expnD -muln_divA ?order_dvdn ?ypn //.
+ by rewrite expgM mem_cycle.
+rewrite !inE mem_cycle -expgM -expnD addnC -maxnE -order_dvdn.
+by rewrite -{1}(part_pnat_id p_x) p_part dvdn_exp2l ?leq_maxr.
+Qed.
+
+Lemma Ohm_dprod A B G : A \x B = G -> 'Ohm_n(A) \x 'Ohm_n(B) = 'Ohm_n(G).
+Proof.
+case/dprodP => [[H K -> ->{A B}]] <- cHK tiHK.
+rewrite dprodEY //; last first.
+- by apply/trivgP; rewrite -tiHK setISS ?Ohm_sub.
+- by rewrite (subset_trans (subset_trans _ cHK)) ?centS ?Ohm_sub.
+apply/eqP; rewrite -(cent_joinEr cHK) eqEsubset join_subG /=.
+rewrite !OhmS ?joing_subl ?joing_subr //= cent_joinEr //= -genM_join genS //.
+apply/subsetP=> _ /setIdP[/imset2P[x y Hx Ky ->] /OhmPredP[p p_pr /eqP]].
+have cxy: commute x y by red; rewrite -(centsP cHK).
+rewrite ?expgMn // -eq_invg_mul => /eqP def_x.
+have ypn1: y ^+ (p ^ n) = 1.
+ by apply/set1P; rewrite -[[set 1]]tiHK inE -{1}def_x groupV !groupX.
+have xpn1: x ^+ (p ^ n) = 1 by rewrite -[x ^+ _]invgK def_x ypn1 invg1.
+by rewrite mem_mulg ?mem_gen // inE (Hx, Ky); apply/OhmPredP; exists p.
+Qed.
+
+Lemma Mho_sub G : 'Mho^n(G) \subset G.
+Proof.
+rewrite gen_subG; apply/subsetP=> _ /imsetP[x /setIdP[Gx _] ->].
+exact: groupX.
+Qed.
+
+Lemma Mho1 : 'Mho^n([1 gT]) = 1. Proof. exact: (trivgP (Mho_sub _)). Qed.
+
+Lemma morphim_Mho rT D G (f : {morphism D >-> rT}) :
+ G \subset D -> f @* 'Mho^n(G) = 'Mho^n(f @* G).
+Proof.
+move=> sGD; have sGnD := subset_trans (Mho_sub G) sGD.
+apply/eqP; rewrite eqEsubset {1}morphim_gen -1?gen_subG // !gen_subG.
+apply/andP; split; apply/subsetP=> y.
+ case/morphimP=> xpn _ /imsetP[x /setIdP[Gx]].
+ set p := pdiv _ => p_x -> -> {xpn y}; have Dx := subsetP sGD x Gx.
+ by rewrite morphX // Mho_p_elt ?morph_p_elt ?mem_morphim.
+case/imsetP=> _ /setIdP[/morphimP[x Dx Gx ->]].
+set p := pdiv _ => p_fx ->{y}; rewrite -(constt_p_elt p_fx) -morph_constt //.
+by rewrite -morphX ?mem_morphim ?Mho_p_elt ?groupX ?p_elt_constt.
+Qed.
+
+Lemma Mho_cont rT G (f : {morphism G >-> rT}) :
+ f @* 'Mho^n(G) \subset 'Mho^n(f @* G).
+Proof. by rewrite morphim_Mho. Qed.
+
+Lemma MhoS H G : H \subset G -> 'Mho^n(H) \subset 'Mho^n(G).
+Proof.
+move=> sHG; apply: genS; apply: imsetS; apply/subsetP=> x.
+by rewrite !inE => /andP[Hx]; rewrite (subsetP sHG).
+Qed.
+
+Lemma MhoE p G : p.-group G -> 'Mho^n(G) = <<[set x ^+ (p ^ n) | x in G]>>.
+Proof.
+move=> pG; apply/eqP; rewrite eqEsubset !gen_subG; apply/andP.
+do [split; apply/subsetP=> xpn; case/imsetP=> x] => [|Gx ->]; last first.
+ by rewrite Mho_p_elt ?(mem_p_elt pG).
+case/setIdP=> Gx _ ->; have [-> | ntx] := eqVneq x 1; first by rewrite expg1n.
+by rewrite (pdiv_p_elt (mem_p_elt pG Gx) ntx) mem_gen //; exact: mem_imset.
+Qed.
+
+Lemma MhoEabelian p G :
+ p.-group G -> abelian G -> 'Mho^n(G) = [set x ^+ (p ^ n) | x in G].
+Proof.
+move=> pG cGG; rewrite (MhoE pG); rewrite gen_set_id //; apply/group_setP.
+split=> [|xn yn]; first by apply/imsetP; exists 1; rewrite ?expg1n.
+case/imsetP=> x Gx ->; case/imsetP=> y Gy ->.
+by rewrite -expgMn; [apply: mem_imset; rewrite groupM | exact: (centsP cGG)].
+Qed.
+
+Lemma trivg_Mho G : 'Mho^n(G) == 1 -> 'Ohm_n(G) == G.
+Proof.
+rewrite -subG1 gen_subG eqEsubset Ohm_sub /= => Gp1.
+rewrite -{1}(Sylow_gen G) genS //; apply/bigcupsP=> P.
+case/SylowP=> p p_pr /and3P[sPG pP _]; apply/subsetP=> x Px.
+have Gx := subsetP sPG x Px; rewrite inE Gx //=.
+rewrite (sameP eqP set1P) (subsetP Gp1) ?mem_gen //; apply: mem_imset.
+by rewrite inE Gx; exact: pgroup_p (mem_p_elt pP Px).
+Qed.
+
+Lemma Mho_p_cycle p x : p.-elt x -> 'Mho^n(<[x]>) = <[x ^+ (p ^ n)]>.
+Proof.
+move=> p_x.
+apply/eqP; rewrite (MhoE p_x) eqEsubset cycle_subG mem_gen; last first.
+ by apply: mem_imset; exact: cycle_id.
+rewrite gen_subG andbT; apply/subsetP=> _ /imsetP[_ /cycleP[k ->] ->].
+by rewrite -expgM mulnC expgM mem_cycle.
+Qed.
+
+Lemma Mho_cprod A B G : A \* B = G -> 'Mho^n(A) \* 'Mho^n(B) = 'Mho^n(G).
+Proof.
+case/cprodP => [[H K -> ->{A B}]] <- cHK; rewrite cprodEY //; last first.
+ by rewrite (subset_trans (subset_trans _ cHK)) ?centS ?Mho_sub.
+apply/eqP; rewrite -(cent_joinEr cHK) eqEsubset join_subG /=.
+rewrite !MhoS ?joing_subl ?joing_subr //= cent_joinEr // -genM_join.
+apply: genS; apply/subsetP=> xypn /imsetP[_ /setIdP[/imset2P[x y Hx Ky ->]]].
+move/constt_p_elt; move: (pdiv _) => p <- ->.
+have cxy: commute x y by red; rewrite -(centsP cHK).
+rewrite consttM // expgMn; last exact: commuteX2.
+by rewrite mem_mulg ?Mho_p_elt ?groupX ?p_elt_constt.
+Qed.
+
+Lemma Mho_dprod A B G : A \x B = G -> 'Mho^n(A) \x 'Mho^n(B) = 'Mho^n(G).
+Proof.
+case/dprodP => [[H K -> ->{A B}]] defG cHK tiHK.
+rewrite dprodEcp; first by apply: Mho_cprod; rewrite cprodE.
+by apply/trivgP; rewrite -tiHK setISS ?Mho_sub.
+Qed.
+
+End Generic.
+
+Canonical Ohm_igFun i := [igFun by Ohm_sub i & Ohm_cont i].
+Canonical Ohm_gFun i := [gFun by Ohm_cont i].
+Canonical Ohm_mgFun i := [mgFun by OhmS i].
+
+Canonical Mho_igFun i := [igFun by Mho_sub i & Mho_cont i].
+Canonical Mho_gFun i := [gFun by Mho_cont i].
+Canonical Mho_mgFun i := [mgFun by MhoS i].
+
+Section char.
+
+Variables (n : nat) (gT rT : finGroupType) (D G : {group gT}).
+
+Lemma Ohm_char : 'Ohm_n(G) \char G. Proof. exact: gFchar. Qed.
+Lemma Ohm_normal : 'Ohm_n(G) <| G. Proof. exact: gFnormal. Qed.
+
+Lemma Mho_char : 'Mho^n(G) \char G. Proof. exact: gFchar. Qed.
+Lemma Mho_normal : 'Mho^n(G) <| G. Proof. exact: gFnormal. Qed.
+
+Lemma morphim_Ohm (f : {morphism D >-> rT}) :
+ G \subset D -> f @* 'Ohm_n(G) \subset 'Ohm_n(f @* G).
+Proof. exact: morphimF. Qed.
+
+Lemma injm_Ohm (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> f @* 'Ohm_n(G) = 'Ohm_n(f @* G).
+Proof. by move=> injf; exact: injmF. Qed.
+
+Lemma isog_Ohm (H : {group rT}) : G \isog H -> 'Ohm_n(G) \isog 'Ohm_n(H).
+Proof. exact: gFisog. Qed.
+
+Lemma isog_Mho (H : {group rT}) : G \isog H -> 'Mho^n(G) \isog 'Mho^n(H).
+Proof. exact: gFisog. Qed.
+
+End char.
+
+Variable gT : finGroupType.
+Implicit Types (pi : nat_pred) (p : nat).
+Implicit Types (A B C : {set gT}) (D G H E : {group gT}).
+
+Lemma Ohm0 G : 'Ohm_0(G) = 1.
+Proof.
+apply/trivgP; rewrite /= gen_subG.
+by apply/subsetP=> x /setIdP[_]; rewrite inE.
+Qed.
+
+Lemma Ohm_leq m n G : m <= n -> 'Ohm_m(G) \subset 'Ohm_n(G).
+Proof.
+move/subnKC <-; rewrite genS //; apply/subsetP=> y.
+by rewrite !inE expnD expgM => /andP[-> /eqP->]; rewrite expg1n /=.
+Qed.
+
+Lemma OhmJ n G x : 'Ohm_n(G :^ x) = 'Ohm_n(G) :^ x.
+Proof.
+rewrite -{1}(setIid G) -(setIidPr (Ohm_sub n G)).
+by rewrite -!morphim_conj injm_Ohm ?injm_conj.
+Qed.
+
+Lemma Mho0 G : 'Mho^0(G) = G.
+Proof.
+apply/eqP; rewrite eqEsubset Mho_sub /=.
+apply/subsetP=> x Gx; rewrite -[x]prod_constt group_prod // => p _.
+exact: Mho_p_elt (groupX _ Gx) (p_elt_constt _ _).
+Qed.
+
+Lemma Mho_leq m n G : m <= n -> 'Mho^n(G) \subset 'Mho^m(G).
+Proof.
+move/subnKC <-; rewrite gen_subG //.
+apply/subsetP=> _ /imsetP[x /setIdP[Gx p_x] ->].
+by rewrite expnD expgM groupX ?(Mho_p_elt _ _ p_x).
+Qed.
+
+Lemma MhoJ n G x : 'Mho^n(G :^ x) = 'Mho^n(G) :^ x.
+Proof.
+by rewrite -{1}(setIid G) -(setIidPr (Mho_sub n G)) -!morphim_conj morphim_Mho.
+Qed.
+
+Lemma extend_cyclic_Mho G p x :
+ p.-group G -> x \in G -> 'Mho^1(G) = <[x ^+ p]> ->
+ forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (p ^ k)]>.
+Proof.
+move=> pG Gx defG1 [//|k _]; have pX := mem_p_elt pG Gx.
+apply/eqP; rewrite eqEsubset cycle_subG (Mho_p_elt _ Gx pX) andbT.
+rewrite (MhoE _ pG) gen_subG; apply/subsetP=> ypk; case/imsetP=> y Gy ->{ypk}.
+have: y ^+ p \in <[x ^+ p]> by rewrite -defG1 (Mho_p_elt 1 _ (mem_p_elt pG Gy)).
+rewrite !expnS /= !expgM => /cycleP[j ->].
+by rewrite -!expgM mulnCA mulnC expgM mem_cycle.
+Qed.
+
+Lemma Ohm1Eprime G : 'Ohm_1(G) = <<[set x in G | prime #[x]]>>.
+Proof.
+rewrite -['Ohm_1(G)](genD1 (group1 _)); congr <<_>>.
+apply/setP=> x; rewrite !inE andbCA -order_dvdn -order_gt1; congr (_ && _).
+apply/andP/idP=> [[p_gt1] | p_pr]; last by rewrite prime_gt1 ?pdiv_id.
+set p := pdiv _ => ox_p; have p_pr: prime p by rewrite pdiv_prime.
+by have [_ dv_p] := primeP p_pr; case/pred2P: (dv_p _ ox_p) p_gt1 => ->.
+Qed.
+
+Lemma abelem_Ohm1 p G : p.-group G -> p.-abelem 'Ohm_1(G) = abelian 'Ohm_1(G).
+Proof.
+move=> pG; rewrite /abelem (pgroupS (Ohm_sub 1 G)) //.
+case abG1: (abelian _) => //=; apply/exponentP=> x.
+by rewrite (OhmEabelian pG abG1); case/LdivP.
+Qed.
+
+Lemma Ohm1_abelem p G : p.-group G -> abelian G -> p.-abelem ('Ohm_1(G)).
+Proof. by move=> pG cGG; rewrite abelem_Ohm1 ?(abelianS (Ohm_sub 1 G)). Qed.
+
+Lemma Ohm1_id p G : p.-abelem G -> 'Ohm_1(G) = G.
+Proof.
+case/and3P=> pG cGG /exponentP Gp.
+apply/eqP; rewrite eqEsubset Ohm_sub (OhmE 1 pG) sub_gen //.
+by apply/subsetP=> x Gx; rewrite !inE Gx Gp /=.
+Qed.
+
+Lemma abelem_Ohm1P p G :
+ abelian G -> p.-group G -> reflect ('Ohm_1(G) = G) (p.-abelem G).
+Proof.
+move=> cGG pG.
+apply: (iffP idP) => [| <-]; [exact: Ohm1_id | exact: Ohm1_abelem].
+Qed.
+
+Lemma TI_Ohm1 G H : H :&: 'Ohm_1(G) = 1 -> H :&: G = 1.
+Proof.
+move=> tiHG1; case: (trivgVpdiv (H :&: G)) => // [[p pr_p]].
+case/Cauchy=> // x /setIP[Hx Gx] ox.
+suffices x1: x \in [1] by rewrite -ox (set1P x1) order1 in pr_p.
+by rewrite -{}tiHG1 inE Hx Ohm1Eprime mem_gen // inE Gx ox.
+Qed.
+
+Lemma Ohm1_eq1 G : ('Ohm_1(G) == 1) = (G :==: 1).
+Proof.
+apply/idP/idP => [/eqP G1_1 | /eqP->]; last by rewrite -subG1 Ohm_sub.
+by rewrite -(setIid G) TI_Ohm1 // G1_1 setIg1.
+Qed.
+
+Lemma meet_Ohm1 G H : G :&: H != 1 -> G :&: 'Ohm_1(H) != 1.
+Proof. by apply: contraNneq => /TI_Ohm1->. Qed.
+
+Lemma Ohm1_cent_max G E p : E \in 'E*_p(G) -> p.-group G -> 'Ohm_1('C_G(E)) = E.
+Proof.
+move=> EpmE pG; have [G1 | ntG]:= eqsVneq G 1.
+ case/pmaxElemP: EpmE; case/pElemP; rewrite G1 => /trivgP-> _ _.
+ by apply/trivgP; rewrite cent1T setIT Ohm_sub.
+have [p_pr _ _] := pgroup_pdiv pG ntG.
+by rewrite (OhmE 1 (pgroupS (subsetIl G _) pG)) (pmaxElem_LdivP _ _) ?genGid.
+Qed.
+
+Lemma Ohm1_cyclic_pgroup_prime p G :
+ cyclic G -> p.-group G -> G :!=: 1 -> #|'Ohm_1(G)| = p.
+Proof.
+move=> cycG pG ntG; set K := 'Ohm_1(G).
+have abelK: p.-abelem K by rewrite Ohm1_abelem ?cyclic_abelian.
+have sKG: K \subset G := Ohm_sub 1 G.
+case/cyclicP: (cyclicS sKG cycG) => x /=; rewrite -/K => defK.
+rewrite defK -orderE (abelem_order_p abelK) //= -/K ?defK ?cycle_id //.
+rewrite -cycle_eq1 -defK -(setIidPr sKG).
+by apply: contraNneq ntG => /TI_Ohm1; rewrite setIid => ->.
+Qed.
+
+Lemma cyclic_pgroup_dprod_trivg p A B C :
+ p.-group C -> cyclic C -> A \x B = C ->
+ A = 1 /\ B = C \/ B = 1 /\ A = C.
+Proof.
+move=> pC cycC; case/cyclicP: cycC pC => x ->{C} pC defC.
+case/dprodP: defC => [] [G H -> ->{A B}] defC _ tiGH; rewrite -defC.
+case: (eqVneq <[x]> 1) => [|ntC].
+ move/trivgP; rewrite -defC mulG_subG => /andP[/trivgP-> _].
+ by rewrite mul1g; left.
+have [pr_p _ _] := pgroup_pdiv pC ntC; pose K := 'Ohm_1(<[x]>).
+have prK : prime #|K| by rewrite (Ohm1_cyclic_pgroup_prime _ pC) ?cycle_cyclic.
+case: (prime_subgroupVti G prK) => [sKG |]; last first.
+ move/TI_Ohm1; rewrite -defC (setIidPl (mulG_subl _ _)) => ->.
+ by left; rewrite mul1g.
+case: (prime_subgroupVti H prK) => [sKH |]; last first.
+ move/TI_Ohm1; rewrite -defC (setIidPl (mulG_subr _ _)) => ->.
+ by right; rewrite mulg1.
+have K1: K :=: 1 by apply/trivgP; rewrite -tiGH subsetI sKG.
+by rewrite K1 cards1 in prK.
+Qed.
+
+Lemma piOhm1 G : \pi('Ohm_1(G)) = \pi(G).
+Proof.
+apply/eq_piP => p; apply/idP/idP; first exact: (piSg (Ohm_sub 1 G)).
+rewrite !mem_primes !cardG_gt0 => /andP[p_pr /Cauchy[] // x Gx oxp].
+by rewrite p_pr -oxp order_dvdG //= Ohm1Eprime mem_gen // inE Gx oxp.
+Qed.
+
+Lemma Ohm1Eexponent p G :
+ prime p -> exponent 'Ohm_1(G) %| p -> 'Ohm_1(G) = 'Ldiv_p(G).
+Proof.
+move=> p_pr expG1p; have pG: p.-group G.
+ apply: sub_in_pnat (pnat_pi (cardG_gt0 G)) => q _.
+ rewrite -piOhm1 mem_primes; case/and3P=> q_pr _; apply: pgroupP q_pr.
+ by rewrite -pnat_exponent (pnat_dvd expG1p) ?pnat_id.
+apply/eqP; rewrite eqEsubset {2}(OhmE 1 pG) subset_gen subsetI Ohm_sub.
+by rewrite sub_LdivT expG1p.
+Qed.
+
+Lemma p_rank_Ohm1 p G : 'r_p('Ohm_1(G)) = 'r_p(G).
+Proof.
+apply/eqP; rewrite eqn_leq p_rankS ?Ohm_sub //.
+apply/bigmax_leqP=> E /setIdP[sEG abelE].
+by rewrite (bigmax_sup E) // inE -{1}(Ohm1_id abelE) OhmS.
+Qed.
+
+Lemma rank_Ohm1 G : 'r('Ohm_1(G)) = 'r(G).
+Proof.
+apply/eqP; rewrite eqn_leq rankS ?Ohm_sub //.
+by have [p _ ->] := rank_witness G; rewrite -p_rank_Ohm1 p_rank_le_rank.
+Qed.
+
+Lemma p_rank_abelian p G : abelian G -> 'r_p(G) = logn p #|'Ohm_1(G)|.
+Proof.
+move=> cGG; have nilG := abelian_nil cGG; case p_pr: (prime p); last first.
+ by apply/eqP; rewrite lognE p_pr eqn0Ngt p_rank_gt0 mem_primes p_pr.
+case/dprodP: (Ohm_dprod 1 (nilpotent_pcoreC p nilG)) => _ <- _ /TI_cardMg->.
+rewrite mulnC logn_Gauss; last first.
+ rewrite prime_coprime // -p'natE // -/(pgroup _ _).
+ exact: pgroupS (Ohm_sub _ _) (pcore_pgroup _ _).
+rewrite -(p_rank_Sylow (nilpotent_pcore_Hall p nilG)) -p_rank_Ohm1.
+rewrite p_rank_abelem // Ohm1_abelem ?pcore_pgroup //.
+exact: abelianS (pcore_sub _ _) cGG.
+Qed.
+
+Lemma rank_abelian_pgroup p G :
+ p.-group G -> abelian G -> 'r(G) = logn p #|'Ohm_1(G)|.
+Proof. by move=> pG cGG; rewrite (rank_pgroup pG) p_rank_abelian. Qed.
+
+End OhmProps.
+
+Section AbelianStructure.
+
+Variable gT : finGroupType.
+Implicit Types (p : nat) (G H K E : {group gT}).
+
+Lemma abelian_splits x G :
+ x \in G -> #[x] = exponent G -> abelian G -> [splits G, over <[x]>].
+Proof.
+move=> Gx ox cGG; apply/splitsP; move: {2}_.+1 (ltnSn #|G|) => n.
+elim: n gT => // n IHn aT in x G Gx ox cGG *; rewrite ltnS => leGn.
+have: <[x]> \subset G by [rewrite cycle_subG]; rewrite subEproper.
+case/predU1P=> [<-|]; first by exists 1%G; rewrite inE -subG1 subsetIr mulg1 /=.
+case/properP=> sxG [y]; elim: {y}_.+1 {-2}y (ltnSn #[y]) => // m IHm y.
+rewrite ltnS => leym Gy x'y; case: (trivgVpdiv <[y]>) => [y1 | [p p_pr p_dv_y]].
+ by rewrite -cycle_subG y1 sub1G in x'y.
+case x_yp: (y ^+ p \in <[x]>); last first.
+ apply: IHm (negbT x_yp); rewrite ?groupX ?(leq_trans _ leym) //.
+ by rewrite orderXdiv // ltn_Pdiv ?prime_gt1.
+have{x_yp} xp_yp: (y ^+ p \in <[x ^+ p]>).
+ have: <[y ^+ p]>%G \in [set <[x ^+ (#[x] %/ #[y ^+ p])]>%G].
+ by rewrite -cycle_sub_group ?order_dvdG // inE cycle_subG x_yp eqxx.
+ rewrite inE -cycle_subG -val_eqE /=; move/eqP->.
+ rewrite cycle_subG orderXdiv // divnA // mulnC ox.
+ by rewrite -muln_divA ?dvdn_exponent ?expgM 1?groupX ?cycle_id.
+have: p <= #[y] by rewrite dvdn_leq.
+rewrite leq_eqVlt; case/predU1P=> [{xp_yp m IHm leym}oy | ltpy]; last first.
+ case/cycleP: xp_yp => k; rewrite -expgM mulnC expgM => def_yp.
+ suffices: #[y * x ^- k] < m.
+ by move/IHm; apply; rewrite groupMr // groupV groupX ?cycle_id.
+ apply: leq_ltn_trans (leq_trans ltpy leym).
+ rewrite dvdn_leq ?prime_gt0 // order_dvdn expgMn.
+ by rewrite expgVn def_yp mulgV.
+ by apply: (centsP cGG); rewrite ?groupV ?groupX.
+pose Y := <[y]>; have nsYG: Y <| G by rewrite -sub_abelian_normal ?cycle_subG.
+have [sYG nYG] := andP nsYG; have nYx := subsetP nYG x Gx.
+have GxY: coset Y x \in G / Y by rewrite mem_morphim.
+have tiYx: Y :&: <[x]> = 1 by rewrite prime_TIg ?indexg1 -?[#|_|]oy ?cycle_subG.
+have: #[coset Y x] = exponent (G / Y).
+ apply/eqP; rewrite eqn_dvd dvdn_exponent //.
+ apply/exponentP=> _ /morphimP[z Nz Gz ->].
+ rewrite -morphX // ((z ^+ _ =P 1) _) ?morph1 //.
+ rewrite orderE -quotient_cycle ?card_quotient ?cycle_subG // -indexgI /=.
+ by rewrite setIC tiYx indexg1 -orderE ox -order_dvdn dvdn_exponent.
+case/IHn => // [||Hq]; first exact: quotient_abelian.
+ apply: leq_trans leGn; rewrite ltn_quotient // cycle_eq1.
+ by apply: contra x'y; move/eqP->; rewrite group1.
+case/complP=> /= ti_x_Hq defGq.
+have: Hq \subset G / Y by rewrite -defGq mulG_subr.
+case/inv_quotientS=> // H defHq sYH sHG; exists H.
+have nYX: <[x]> \subset 'N(Y) by rewrite cycle_subG.
+rewrite inE -subG1 eqEsubset mul_subG //= -tiYx subsetI subsetIl andbT.
+rewrite -{2}(mulSGid sYH) mulgA (normC nYX) -mulgA -quotientSK ?quotientMl //.
+rewrite -quotient_sub1 ?(subset_trans (subsetIl _ _)) // quotientIG //= -/Y.
+by rewrite -defHq quotient_cycle // ti_x_Hq defGq !subxx.
+Qed.
+
+Lemma abelem_splits p G H : p.-abelem G -> H \subset G -> [splits G, over H].
+Proof.
+elim: {G}_.+1 {-2}G H (ltnSn #|G|) => // m IHm G H.
+rewrite ltnS => leGm abelG sHG; case: (eqsVneq H 1) => [-> | ].
+ by apply/splitsP; exists G; rewrite inE mul1g -subG1 subsetIl /=.
+case/trivgPn=> x Hx ntx; have Gx := subsetP sHG x Hx.
+have [_ cGG eGp] := and3P abelG.
+have ox: #[x] = exponent G.
+ by apply/eqP; rewrite eqn_dvd dvdn_exponent // (abelem_order_p abelG).
+case/splitsP: (abelian_splits Gx ox cGG) => K; case/complP=> tixK defG.
+have sKG: K \subset G by rewrite -defG mulG_subr.
+have ltKm: #|K| < m.
+ rewrite (leq_trans _ leGm) ?proper_card //; apply/properP; split=> //.
+ exists x => //; apply: contra ntx => Kx; rewrite -cycle_eq1 -subG1 -tixK.
+ by rewrite subsetI subxx cycle_subG.
+case/splitsP: (IHm _ _ ltKm (abelemS sKG abelG) (subsetIr H K)) => L.
+case/complP=> tiHKL defK; apply/splitsP; exists L; rewrite inE.
+rewrite -subG1 -tiHKL -setIA setIS; last by rewrite subsetI -defK mulG_subr /=.
+by rewrite -(setIidPr sHG) -defG -group_modl ?cycle_subG //= setIC -mulgA defK.
+Qed.
+
+Fact abelian_type_subproof G :
+ {H : {group gT} & abelian G -> {x | #[x] = exponent G & <[x]> \x H = G}}.
+Proof.
+case cGG: (abelian G); last by exists G.
+have [x Gx ox] := exponent_witness (abelian_nil cGG).
+case/splitsP/ex_mingroup: (abelian_splits Gx (esym ox) cGG) => H.
+case/mingroupp/complP=> tixH defG; exists H => _.
+exists x; rewrite ?dprodE // (sub_abelian_cent2 cGG) ?cycle_subG //.
+by rewrite -defG mulG_subr.
+Qed.
+
+Fixpoint abelian_type_rec n G :=
+ if n is n'.+1 then if abelian G && (G :!=: 1) then
+ exponent G :: abelian_type_rec n' (tag (abelian_type_subproof G))
+ else [::] else [::].
+
+Definition abelian_type (A : {set gT}) := abelian_type_rec #|A| <<A>>.
+
+Lemma abelian_type_dvdn_sorted A : sorted [rel m n | n %| m] (abelian_type A).
+Proof.
+set R := SimplRel _; pose G := <<A>>%G.
+suffices: path R (exponent G) (abelian_type A) by case: (_ A) => // m t /andP[].
+rewrite /abelian_type -/G; elim: {A}#|A| G {2 3}G (subxx G) => // n IHn G M sGM.
+simpl; case: ifP => //= /andP[cGG ntG]; rewrite exponentS ?IHn //=.
+case: (abelian_type_subproof G) => H /= [//| x _] /dprodP[_ /= <- _ _].
+exact: mulG_subr.
+Qed.
+
+Lemma abelian_type_gt1 A : all [pred m | m > 1] (abelian_type A).
+Proof.
+rewrite /abelian_type; elim: {A}#|A| <<A>>%G => //= n IHn G.
+case: ifP => //= /andP[_ ntG]; rewrite {n}IHn.
+by rewrite ltn_neqAle exponent_gt0 eq_sym -dvdn1 -trivg_exponent ntG.
+Qed.
+
+Lemma abelian_type_sorted A : sorted geq (abelian_type A).
+Proof.
+have:= abelian_type_dvdn_sorted A; have:= abelian_type_gt1 A.
+case: (abelian_type A) => //= m t; elim: t m => //= n t IHt m /andP[].
+by move/ltnW=> m_gt0 t_gt1 /andP[n_dv_m /IHt->]; rewrite // dvdn_leq.
+Qed.
+
+Theorem abelian_structure G :
+ abelian G ->
+ {b | \big[dprod/1]_(x <- b) <[x]> = G & map order b = abelian_type G}.
+Proof.
+rewrite /abelian_type genGidG.
+elim: {G}#|G| {-2 5}G (leqnn #|G|) => /= [|n IHn] G leGn cGG.
+ by rewrite leqNgt cardG_gt0 in leGn.
+rewrite {1}cGG /=; case: ifP => [ntG|/eqP->]; last first.
+ by exists [::]; rewrite ?big_nil.
+case: (abelian_type_subproof G) => H /= [//|x ox xdefG]; rewrite -ox.
+have [_ defG cxH tixH] := dprodP xdefG.
+have sHG: H \subset G by rewrite -defG mulG_subr.
+case/IHn: (abelianS sHG cGG) => [|b defH <-].
+ rewrite -ltnS (leq_trans _ leGn) // -defG TI_cardMg // -orderE.
+ rewrite ltn_Pmull ?cardG_gt0 // ltn_neqAle order_gt0 eq_sym -dvdn1.
+ by rewrite ox -trivg_exponent ntG.
+by exists (x :: b); rewrite // big_cons defH xdefG.
+Qed.
+
+Lemma count_logn_dprod_cycle p n b G :
+ \big[dprod/1]_(x <- b) <[x]> = G ->
+ count [pred x | logn p #[x] > n] b = logn p #|'Ohm_n.+1(G) : 'Ohm_n(G)|.
+Proof.
+have sOn1 := @Ohm_leq gT _ _ _ (leqnSn n).
+pose lnO i (A : {set gT}) := logn p #|'Ohm_i(A)|.
+have lnO_le H: lnO n H <= lnO n.+1 H.
+ by rewrite dvdn_leq_log ?cardG_gt0 // cardSg ?sOn1.
+have lnOx i A B H: A \x B = H -> lnO i A + lnO i B = lnO i H.
+ move=> defH; case/dprodP: defH (defH) => {A B}[[A B -> ->]] _ _ _ defH.
+ rewrite /lnO; case/dprodP: (Ohm_dprod i defH) => _ <- _ tiOAB.
+ by rewrite TI_cardMg ?lognM.
+rewrite -divgS //= logn_div ?cardSg //= -/(lnO _ _) -/(lnO _ _).
+elim: b G => [_ <-|x b IHb G] /=.
+ by rewrite big_nil /lnO !(trivgP (Ohm_sub _ _)) subnn.
+rewrite /= big_cons => defG; rewrite -!(lnOx _ _ _ _ defG) subnDA.
+case/dprodP: defG => [[_ H _ defH] _ _ _] {G}; rewrite defH (IHb _ defH).
+symmetry; do 2!rewrite addnC -addnBA ?lnO_le //; congr (_ + _).
+pose y := x.`_p; have p_y: p.-elt y by rewrite p_elt_constt.
+have{lnOx} lnOy i: lnO i <[x]> = lnO i <[y]>.
+ have cXX := cycle_abelian x.
+ have co_yx': coprime #[y] #[x.`_p^'] by rewrite !order_constt coprime_partC.
+ have defX: <[y]> \x <[x.`_p^']> = <[x]>.
+ rewrite dprodE ?coprime_TIg //.
+ by rewrite -cycleM ?consttC //; apply: (centsP cXX); exact: mem_cycle.
+ by apply: (sub_abelian_cent2 cXX); rewrite cycle_subG mem_cycle.
+ rewrite -(lnOx i _ _ _ defX) addnC {1}/lnO lognE.
+ case: and3P => // [[p_pr _ /idPn[]]]; rewrite -p'natE //.
+ exact: pgroupS (Ohm_sub _ _) (p_elt_constt _ _).
+rewrite -logn_part -order_constt -/y !{}lnOy /lnO !(Ohm_p_cycle _ p_y).
+case: leqP => [| lt_n_y].
+ by rewrite -subn_eq0 -addn1 subnDA => /eqP->; rewrite subnn.
+rewrite -!orderE -(subSS n) subSn // expnSr expgM.
+have p_pr: prime p by move: lt_n_y; rewrite lognE; case: prime.
+set m := (p ^ _)%N; have m_gt0: m > 0 by rewrite expn_gt0 prime_gt0.
+suffices p_ym: p %| #[y ^+ m].
+ rewrite -logn_div ?orderXdvd // (orderXdiv p_ym) divnA // mulKn //.
+ by rewrite logn_prime ?eqxx.
+rewrite orderXdiv ?pfactor_dvdn ?leq_subr // -(dvdn_pmul2r m_gt0).
+by rewrite -expnS -subSn // subSS divnK pfactor_dvdn ?leq_subr.
+Qed.
+
+Lemma perm_eq_abelian_type p b G :
+ p.-group G -> \big[dprod/1]_(x <- b) <[x]> = G -> 1 \notin b ->
+ perm_eq (map order b) (abelian_type G).
+Proof.
+move: b => b1 pG defG1 ntb1.
+have cGG: abelian G.
+ elim: (b1) {pG}G defG1 => [_ <-|x b IHb G]; first by rewrite big_nil abelian1.
+ rewrite big_cons; case/dprodP=> [[_ H _ defH]] <-; rewrite defH => cxH _.
+ by rewrite abelianM cycle_abelian IHb.
+have p_bG b: \big[dprod/1]_(x <- b) <[x]> = G -> all (p_elt p) b.
+ elim: b {defG1 cGG}G pG => //= x b IHb G pG; rewrite big_cons.
+ case/dprodP=> [[_ H _ defH]]; rewrite defH andbC => defG _ _.
+ by rewrite -defG pgroupM in pG; case/andP: pG => p_x /IHb->.
+have [b2 defG2 def_t] := abelian_structure cGG.
+have ntb2: 1 \notin b2.
+ apply: contraL (abelian_type_gt1 G) => b2_1.
+ rewrite -def_t -has_predC has_map.
+ by apply/hasP; exists 1; rewrite //= order1.
+rewrite -{}def_t; apply/allP=> m; rewrite -map_cat => /mapP[x b_x def_m].
+have{ntb1 ntb2} ntx: x != 1.
+ by apply: contraL b_x; move/eqP->; rewrite mem_cat negb_or ntb1 ntb2.
+have p_x: p.-elt x by apply: allP (x) b_x; rewrite all_cat !p_bG.
+rewrite -cycle_eq1 in ntx; have [p_pr _ [k ox]] := pgroup_pdiv p_x ntx.
+apply/eqnP; rewrite {m}def_m orderE ox !count_map.
+pose cnt_p k := count [pred x : gT | logn p #[x] > k].
+have cnt_b b: \big[dprod/1]_(x <- b) <[x]> = G ->
+ count [pred x | #[x] == p ^ k.+1]%N b = cnt_p k b - cnt_p k.+1 b.
+- move/p_bG; elim: b => //= _ b IHb /andP[/p_natP[j ->] /IHb-> {IHb}].
+ rewrite eqn_leq !leq_exp2l ?prime_gt1 // -eqn_leq pfactorK // leqNgt.
+ case: ltngtP => // _ {j}; rewrite subSn // add0n; elim: b => //= y b IHb.
+ by rewrite leq_add // ltn_neqAle; case: (~~ _).
+by rewrite !cnt_b // /cnt_p !(@count_logn_dprod_cycle _ _ _ G).
+Qed.
+
+Lemma size_abelian_type G : abelian G -> size (abelian_type G) = 'r(G).
+Proof.
+move=> cGG; have [b defG def_t] := abelian_structure cGG.
+apply/eqP; rewrite -def_t size_map eqn_leq andbC; apply/andP; split.
+ have [p p_pr ->] := rank_witness G; rewrite p_rank_abelian //.
+ by rewrite -indexg1 -(Ohm0 G) -(count_logn_dprod_cycle _ _ defG) count_size.
+case/lastP def_b: b => // [b' x]; pose p := pdiv #[x].
+have p_pr: prime p.
+ have:= abelian_type_gt1 G; rewrite -def_t def_b map_rcons -cats1 all_cat.
+ by rewrite /= andbT => /andP[_]; exact: pdiv_prime.
+suffices: all [pred y | logn p #[y] > 0] b.
+ rewrite all_count (count_logn_dprod_cycle _ _ defG) -def_b; move/eqP <-.
+ by rewrite Ohm0 indexg1 -p_rank_abelian ?p_rank_le_rank.
+apply/allP=> y; rewrite def_b mem_rcons inE /= => b_y.
+rewrite lognE p_pr order_gt0 (dvdn_trans (pdiv_dvd _)) //.
+case/predU1P: b_y => [-> // | b'_y].
+have:= abelian_type_dvdn_sorted G; rewrite -def_t def_b.
+case/splitPr: b'_y => b1 b2; rewrite -cat_rcons rcons_cat map_cat !map_rcons.
+rewrite headI /= cat_path -(last_cons 2) -headI last_rcons.
+case/andP=> _ /order_path_min min_y.
+apply: (allP (min_y _)) => [? ? ? ? dv|]; first exact: (dvdn_trans dv).
+by rewrite mem_rcons mem_head.
+Qed.
+
+Lemma mul_card_Ohm_Mho_abelian n G :
+ abelian G -> (#|'Ohm_n(G)| * #|'Mho^n(G)|)%N = #|G|.
+Proof.
+case/abelian_structure => b defG _.
+elim: b G defG => [_ <-|x b IHb G].
+ by rewrite !big_nil (trivgP (Ohm_sub _ _)) (trivgP (Mho_sub _ _)) !cards1.
+rewrite big_cons => defG; rewrite -(dprod_card defG).
+rewrite -(dprod_card (Ohm_dprod n defG)) -(dprod_card (Mho_dprod n defG)) /=.
+rewrite mulnCA -!mulnA mulnCA mulnA; case/dprodP: defG => [[_ H _ defH] _ _ _].
+rewrite defH {b G defH IHb}(IHb H defH); congr (_ * _)%N => {H}.
+elim: {x}_.+1 {-2}x (ltnSn #[x]) => // m IHm x; rewrite ltnS => lexm.
+case p_x: (p_group <[x]>); last first.
+ case: (eqVneq x 1) p_x => [-> |]; first by rewrite cycle1 p_group1.
+ rewrite -order_gt1 /p_group -orderE; set p := pdiv _ => ntx p'x.
+ have def_x: <[x.`_p]> \x <[x.`_p^']> = <[x]>.
+ have ?: coprime #[x.`_p] #[x.`_p^'] by rewrite !order_constt coprime_partC.
+ have ?: commute x.`_p x.`_p^' by exact: commuteX2.
+ rewrite dprodE ?coprime_TIg -?cycleM ?consttC //.
+ by rewrite cent_cycle cycle_subG; exact/cent1P.
+ rewrite -(dprod_card (Ohm_dprod n def_x)) -(dprod_card (Mho_dprod n def_x)).
+ rewrite mulnCA -mulnA mulnCA mulnA.
+ rewrite !{}IHm ?(dprod_card def_x) ?(leq_trans _ lexm) {m lexm}//.
+ rewrite /order -(dprod_card def_x) -!orderE !order_constt ltn_Pmull //.
+ rewrite p_part -(expn0 p) ltn_exp2l 1?lognE ?prime_gt1 ?pdiv_prime //.
+ by rewrite order_gt0 pdiv_dvd.
+ rewrite proper_card // properEneq cycle_subG mem_cycle andbT.
+ by apply: contra (negbT p'x); move/eqP <-; exact: p_elt_constt.
+case/p_groupP: p_x => p p_pr p_x.
+rewrite (Ohm_p_cycle n p_x) (Mho_p_cycle n p_x) -!orderE.
+set k := logn p #[x]; have ox: #[x] = (p ^ k)%N by rewrite -card_pgroup.
+case: (leqP k n) => [le_k_n | lt_n_k].
+ rewrite -(subnKC le_k_n) subnDA subnn expg1 expnD expgM -ox.
+ by rewrite expg_order expg1n order1 muln1.
+rewrite !orderXgcd ox -{-3}(subnKC (ltnW lt_n_k)) expnD.
+rewrite gcdnC gcdnMl gcdnC gcdnMr.
+by rewrite mulnK ?mulKn ?expn_gt0 ?prime_gt0.
+Qed.
+
+Lemma grank_abelian G : abelian G -> 'm(G) = 'r(G).
+Proof.
+move=> cGG; apply/eqP; rewrite eqn_leq; apply/andP; split.
+ rewrite -size_abelian_type //; case/abelian_structure: cGG => b defG <-.
+ suffices <-: <<[set x in b]>> = G.
+ by rewrite (leq_trans (grank_min _)) // size_map cardsE card_size.
+ rewrite -{G defG}(bigdprodWY defG).
+ elim: b => [|x b IHb]; first by rewrite big_nil gen0.
+ by rewrite big_cons -joingE -joing_idr -IHb joing_idl joing_idr set_cons.
+have [p p_pr ->] := rank_witness G; pose K := 'Mho^1(G).
+have ->: 'r_p(G) = logn p #|G / K|.
+ rewrite p_rank_abelian // card_quotient /= ?gFnorm // -divgS ?Mho_sub //.
+ by rewrite -(mul_card_Ohm_Mho_abelian 1 cGG) mulnK ?cardG_gt0.
+case: (grank_witness G) => B genB <-; rewrite -genB.
+have: <<B>> \subset G by rewrite genB.
+elim: {B genB}_.+1 {-2}B (ltnSn #|B|) => // m IHm B; rewrite ltnS.
+case: (set_0Vmem B) => [-> | [x Bx]].
+ by rewrite gen0 quotient1 cards1 logn1.
+rewrite (cardsD1 x) Bx -{2 3}(setD1K Bx); set B' := B :\ x => ltB'm.
+rewrite -joingE -joing_idl -joing_idr -/<[x]> join_subG => /andP[Gx sB'G].
+rewrite cent_joinEl ?(sub_abelian_cent2 cGG) //.
+have nKx: x \in 'N(K) by rewrite -cycle_subG (subset_trans Gx) ?gFnorm.
+rewrite quotientMl ?cycle_subG // quotient_cycle //= -/K.
+have le_Kxp_1: logn p #[coset K x] <= 1.
+ rewrite -(dvdn_Pexp2l _ _ (prime_gt1 p_pr)) -p_part -order_constt.
+ rewrite order_dvdn -morph_constt // -morphX ?groupX //= coset_id //.
+ by rewrite Mho_p_elt ?p_elt_constt ?groupX -?cycle_subG.
+apply: leq_trans (leq_add le_Kxp_1 (IHm _ ltB'm sB'G)).
+by rewrite -lognM ?dvdn_leq_log ?muln_gt0 ?cardG_gt0 // mul_cardG dvdn_mulr.
+Qed.
+
+Lemma rank_cycle (x : gT) : 'r(<[x]>) = (x != 1).
+Proof.
+have [->|ntx] := altP (x =P 1); first by rewrite cycle1 rank1.
+apply/eqP; rewrite eqn_leq rank_gt0 cycle_eq1 ntx andbT.
+by rewrite -grank_abelian ?cycle_abelian //= -(cards1 x) grank_min.
+Qed.
+
+Lemma abelian_rank1_cyclic G : abelian G -> cyclic G = ('r(G) <= 1).
+Proof.
+move=> cGG; have [b defG atypG] := abelian_structure cGG.
+apply/idP/idP; first by case/cyclicP=> x ->; rewrite rank_cycle leq_b1.
+rewrite -size_abelian_type // -{}atypG -{}defG unlock.
+by case: b => [|x []] //= _; rewrite ?cyclic1 // dprodg1 cycle_cyclic.
+Qed.
+
+Definition homocyclic A := abelian A && constant (abelian_type A).
+
+Lemma homocyclic_Ohm_Mho n p G :
+ p.-group G -> homocyclic G -> 'Ohm_n(G) = 'Mho^(logn p (exponent G) - n)(G).
+Proof.
+move=> pG /andP[cGG homoG]; set e := exponent G.
+have{pG} p_e: p.-nat e by apply: pnat_dvd pG; exact: exponent_dvdn.
+have{homoG}: all (pred1 e) (abelian_type G).
+ move: homoG; rewrite /abelian_type -(prednK (cardG_gt0 G)) /=.
+ by case: (_ && _) (tag _); rewrite //= genGid eqxx.
+have{cGG} [b defG <-] := abelian_structure cGG.
+move: e => e in p_e *; elim: b => /= [|x b IHb] in G defG *.
+ by rewrite -defG big_nil (trivgP (Ohm_sub _ _)) (trivgP (Mho_sub _ _)).
+case/andP=> /eqP ox e_b; rewrite big_cons in defG.
+rewrite -(Ohm_dprod _ defG) -(Mho_dprod _ defG).
+case/dprodP: defG => [[_ H _ defH] _ _ _]; rewrite defH IHb //; congr (_ \x _).
+by rewrite -ox in p_e *; rewrite (Ohm_p_cycle _ p_e) (Mho_p_cycle _ p_e).
+Qed.
+
+Lemma Ohm_Mho_homocyclic (n p : nat) G :
+ abelian G -> p.-group G -> 0 < n < logn p (exponent G) ->
+ 'Ohm_n(G) = 'Mho^(logn p (exponent G) - n)(G) -> homocyclic G.
+Proof.
+set e := exponent G => cGG pG /andP[n_gt0 n_lte] eq_Ohm_Mho.
+suffices: all (pred1 e) (abelian_type G).
+ by rewrite /homocyclic cGG; exact: all_pred1_constant.
+case/abelian_structure: cGG (abelian_type_gt1 G) => b defG <-.
+elim: b {-3}G defG (subxx G) eq_Ohm_Mho => //= x b IHb H.
+rewrite big_cons => defG; case/dprodP: defG (defG) => [[_ K _ defK]].
+rewrite defK => defHm cxK; rewrite setIC; move/trivgP=> tiKx defHd.
+rewrite -{1}defHm {defHm} mulG_subG cycle_subG ltnNge -trivg_card_le1.
+case/andP=> Gx sKG; rewrite -(Mho_dprod _ defHd) => /esym defMho /andP[ntx ntb].
+have{defHd} defOhm := Ohm_dprod n defHd.
+apply/andP; split; last first.
+ apply: (IHb K) => //; have:= dprod_modr defMho (Mho_sub _ _).
+ rewrite -(dprod_modr defOhm (Ohm_sub _ _)).
+ rewrite !(trivgP (subset_trans (setIS _ _) tiKx)) ?Ohm_sub ?Mho_sub //.
+ by rewrite !dprod1g.
+have:= dprod_modl defMho (Mho_sub _ _).
+rewrite -(dprod_modl defOhm (Ohm_sub _ _)) .
+rewrite !(trivgP (subset_trans (setSI _ _) tiKx)) ?Ohm_sub ?Mho_sub //.
+move/eqP; rewrite eqEcard => /andP[_].
+have p_x: p.-elt x := mem_p_elt pG Gx.
+have [p_pr p_dv_x _] := pgroup_pdiv p_x ntx.
+rewrite !dprodg1 (Ohm_p_cycle _ p_x) (Mho_p_cycle _ p_x) -!orderE.
+rewrite orderXdiv ?leq_divLR ?pfactor_dvdn ?leq_subr //.
+rewrite orderXgcd divn_mulAC ?dvdn_gcdl // leq_divRL ?gcdn_gt0 ?order_gt0 //.
+rewrite leq_pmul2l //; apply: contraLR.
+rewrite eqn_dvd dvdn_exponent //= -ltnNge => lt_x_e.
+rewrite (leq_trans (ltn_Pmull (prime_gt1 p_pr) _)) ?expn_gt0 ?prime_gt0 //.
+rewrite -expnS dvdn_leq // ?gcdn_gt0 ?order_gt0 // dvdn_gcd.
+rewrite pfactor_dvdn // dvdn_exp2l.
+ by rewrite -{2}[logn p _]subn0 ltn_sub2l // lognE p_pr order_gt0 p_dv_x.
+rewrite ltn_sub2r // ltnNge -(dvdn_Pexp2l _ _ (prime_gt1 p_pr)) -!p_part.
+by rewrite !part_pnat_id // (pnat_dvd (exponent_dvdn G)).
+Qed.
+
+Lemma abelem_homocyclic p G : p.-abelem G -> homocyclic G.
+Proof.
+move=> abelG; have [_ cGG _] := and3P abelG.
+rewrite /homocyclic cGG (@all_pred1_constant _ p) //.
+case/abelian_structure: cGG (abelian_type_gt1 G) => b defG <- => b_gt1.
+apply/allP=> _ /mapP[x b_x ->] /=; rewrite (abelem_order_p abelG) //.
+ rewrite -cycle_subG -(bigdprodWY defG) ?sub_gen //.
+ by rewrite bigcup_seq (bigcup_sup x).
+by rewrite -order_gt1 [_ > 1](allP b_gt1) ?map_f.
+Qed.
+
+Lemma homocyclic1 : homocyclic [1 gT].
+Proof. exact: abelem_homocyclic (abelem1 _ 2). Qed.
+
+Lemma Ohm1_homocyclicP p G : p.-group G -> abelian G ->
+ reflect ('Ohm_1(G) = 'Mho^(logn p (exponent G)).-1(G)) (homocyclic G).
+Proof.
+move=> pG cGG; set e := logn p (exponent G); rewrite -subn1.
+apply: (iffP idP) => [homoG | ]; first exact: homocyclic_Ohm_Mho.
+case: (ltnP 1 e) => [lt1e | ]; first exact: Ohm_Mho_homocyclic.
+rewrite -subn_eq0 => /eqP->; rewrite Mho0 => <-.
+exact: abelem_homocyclic (Ohm1_abelem pG cGG).
+Qed.
+
+Lemma abelian_type_homocyclic G :
+ homocyclic G -> abelian_type G = nseq 'r(G) (exponent G).
+Proof.
+case/andP=> cGG; rewrite -size_abelian_type // /abelian_type.
+rewrite -(prednK (cardG_gt0 G)) /=; case: andP => //= _; move: (tag _) => H.
+by move/all_pred1P->; rewrite genGid size_nseq.
+Qed.
+
+Lemma abelian_type_abelem p G : p.-abelem G -> abelian_type G = nseq 'r(G) p.
+Proof.
+move=> abelG; rewrite (abelian_type_homocyclic (abelem_homocyclic abelG)).
+case: (eqVneq G 1%G) => [-> | ntG]; first by rewrite rank1.
+congr nseq; apply/eqP; rewrite eqn_dvd; have [pG _ ->] := and3P abelG.
+have [p_pr] := pgroup_pdiv pG ntG; case/Cauchy=> // x Gx <- _.
+exact: dvdn_exponent.
+Qed.
+
+Lemma max_card_abelian G :
+ abelian G -> #|G| <= exponent G ^ 'r(G) ?= iff homocyclic G.
+Proof.
+move=> cGG; have [b defG def_tG] := abelian_structure cGG.
+have Gb: all (mem G) b.
+ apply/allP=> x b_x; rewrite -(bigdprodWY defG); have [b1 b2] := splitPr b_x.
+ by rewrite big_cat big_cons /= mem_gen // setUCA inE cycle_id.
+have ->: homocyclic G = all (pred1 (exponent G)) (abelian_type G).
+ rewrite /homocyclic cGG /abelian_type; case: #|G| => //= n.
+ by move: (_ (tag _)) => t; case: ifP => //= _; rewrite genGid eqxx.
+rewrite -size_abelian_type // -{}def_tG -{defG}(bigdprod_card defG) size_map.
+rewrite unlock; elim: b Gb => //= x b IHb; case/andP=> Gx Gb.
+have eGgt0: exponent G > 0 := exponent_gt0 G.
+have le_x_G: #[x] <= exponent G by rewrite dvdn_leq ?dvdn_exponent.
+have:= leqif_mul (leqif_eq le_x_G) (IHb Gb).
+by rewrite -expnS expn_eq0 eqn0Ngt eGgt0.
+Qed.
+
+Lemma card_homocyclic G : homocyclic G -> #|G| = (exponent G ^ 'r(G))%N.
+Proof.
+by move=> homG; have [cGG _] := andP homG; apply/eqP; rewrite max_card_abelian.
+Qed.
+
+Lemma abelian_type_dprod_homocyclic p K H G :
+ K \x H = G -> p.-group G -> homocyclic G ->
+ abelian_type K = nseq 'r(K) (exponent G)
+ /\ abelian_type H = nseq 'r(H) (exponent G).
+Proof.
+move=> defG pG homG; have [cGG _] := andP homG.
+have /mulG_sub[sKG sHG]: K * H = G by case/dprodP: defG.
+have [cKK cHH] := (abelianS sKG cGG, abelianS sHG cGG).
+suffices: all (pred1 (exponent G)) (abelian_type K ++ abelian_type H).
+ rewrite all_cat => /andP[/all_pred1P-> /all_pred1P->].
+ by rewrite !size_abelian_type.
+suffices def_atG: abelian_type K ++ abelian_type H =i abelian_type G.
+ rewrite (eq_all_r def_atG); apply/all_pred1P.
+ by rewrite size_abelian_type // -abelian_type_homocyclic.
+have [bK defK atK] := abelian_structure cKK.
+have [bH defH atH] := abelian_structure cHH.
+apply: perm_eq_mem; rewrite -atK -atH -map_cat.
+apply: (perm_eq_abelian_type pG); first by rewrite big_cat defK defH.
+have: all [pred m | m > 1] (map order (bK ++ bH)).
+ by rewrite map_cat all_cat atK atH !abelian_type_gt1.
+by rewrite all_map (eq_all (@order_gt1 _)) all_predC has_pred1.
+Qed.
+
+Lemma dprod_homocyclic p K H G :
+ K \x H = G -> p.-group G -> homocyclic G -> homocyclic K /\ homocyclic H.
+Proof.
+move=> defG pG homG; have [cGG _] := andP homG.
+have /mulG_sub[sKG sHG]: K * H = G by case/dprodP: defG.
+have [abtK abtH] := abelian_type_dprod_homocyclic defG pG homG.
+by rewrite /homocyclic !(abelianS _ cGG) // abtK abtH !constant_nseq.
+Qed.
+
+Lemma exponent_dprod_homocyclic p K H G :
+ K \x H = G -> p.-group G -> homocyclic G -> K :!=: 1 ->
+ exponent K = exponent G.
+Proof.
+move=> defG pG homG ntK; have [homK _] := dprod_homocyclic defG pG homG.
+have [] := abelian_type_dprod_homocyclic defG pG homG.
+by rewrite abelian_type_homocyclic // -['r(K)]prednK ?rank_gt0 => [[]|].
+Qed.
+
+End AbelianStructure.
+
+Arguments Scope abelian_type [_ group_scope].
+Arguments Scope homocyclic [_ group_scope].
+Prenex Implicits abelian_type homocyclic.
+
+Section IsogAbelian.
+
+Variables aT rT : finGroupType.
+Implicit Type (gT : finGroupType) (D G : {group aT}) (H : {group rT}).
+
+Lemma isog_abelian_type G H : isog G H -> abelian_type G = abelian_type H.
+Proof.
+pose lnO p n gT (A : {set gT}) := logn p #|'Ohm_n.+1(A) : 'Ohm_n(A)|.
+pose lni i p gT (A : {set gT}) := \max_(e < logn p #|A| | i < lnO p e _ A) e.+1.
+suffices{G} nth_abty gT (G : {group gT}) i:
+ abelian G -> i < size (abelian_type G) ->
+ nth 1%N (abelian_type G) i = (\prod_(p < #|G|.+1) p ^ lni i p _ G)%N.
+- move=> isoGH; case cGG: (abelian G); last first.
+ rewrite /abelian_type -(prednK (cardG_gt0 G)) -(prednK (cardG_gt0 H)) /=.
+ by rewrite {1}(genGid G) {1}(genGid H) -(isog_abelian isoGH) cGG.
+ have cHH: abelian H by rewrite -(isog_abelian isoGH).
+ have eq_sz: size (abelian_type G) = size (abelian_type H).
+ by rewrite !size_abelian_type ?(isog_rank isoGH).
+ apply: (@eq_from_nth _ 1%N) => // i lt_i_G; rewrite !nth_abty // -?eq_sz //.
+ rewrite /lni (card_isog isoGH); apply: eq_bigr => p _; congr (p ^ _)%N.
+ apply: eq_bigl => e; rewrite /lnO -!divgS ?(Ohm_leq _ (leqnSn _)) //=.
+ by have:= card_isog (gFisog _ isoGH) => /= eqF; rewrite !eqF.
+move=> cGG.
+have (p): path leq 0 (map (logn p) (rev (abelian_type G))).
+ move: (abelian_type_gt1 G) (abelian_type_dvdn_sorted G).
+ case: abelian_type => //= m t; rewrite rev_cons map_rcons.
+ elim: t m => //= n t IHt m /andP[/ltnW m_gt0 nt_gt1].
+ rewrite -cats1 cat_path rev_cons map_rcons last_rcons /=.
+ by case/andP=> /dvdn_leq_log-> // /IHt->.
+have{cGG} [b defG <- b_sorted] := abelian_structure cGG.
+rewrite size_map => ltib; rewrite (nth_map 1 _ _ ltib); set x := nth 1 b i.
+have Gx: x \in G.
+ have: x \in b by rewrite mem_nth.
+ rewrite -(bigdprodWY defG); case/splitPr=> bl br.
+ by rewrite mem_gen // big_cat big_cons !inE cycle_id orbT.
+have lexG: #[x] <= #|G| by rewrite dvdn_leq ?order_dvdG.
+rewrite -[#[x]]partn_pi // (widen_partn _ lexG) big_mkord big_mkcond.
+apply: eq_bigr => p _; transitivity (p ^ logn p #[x])%N.
+ by rewrite -logn_gt0; case: posnP => // ->.
+suffices lti_lnO e: (i < lnO p e _ G) = (e < logn p #[x]).
+ congr (p ^ _)%N; apply/eqP; rewrite eqn_leq andbC; apply/andP; split.
+ by apply/bigmax_leqP=> e; rewrite lti_lnO.
+ case: (posnP (logn p #[x])) => [-> // | logx_gt0].
+ have lexpG: (logn p #[x]).-1 < logn p #|G|.
+ by rewrite prednK // dvdn_leq_log ?order_dvdG.
+ by rewrite (@bigmax_sup _ (Ordinal lexpG)) ?(prednK, lti_lnO).
+rewrite /lnO -(count_logn_dprod_cycle _ _ defG).
+case: (ltnP e _) (b_sorted p) => [lt_e_x | le_x_e].
+ rewrite -(cat_take_drop i.+1 b) -map_rev rev_cat !map_cat cat_path.
+ case/andP=> _ ordb; rewrite count_cat ((count _ _ =P i.+1) _) ?leq_addr //.
+ rewrite -{2}(size_takel ltib) -all_count.
+ move: ordb; rewrite (take_nth 1 ltib) -/x rev_rcons all_rcons /= lt_e_x.
+ case/andP=> _ /=; move/(order_path_min leq_trans); apply: contraLR.
+ rewrite -!has_predC !has_map; case/hasP=> y b_y /= le_y_e; apply/hasP.
+ by exists y; rewrite ?mem_rev //=; apply: contra le_y_e; exact: leq_trans.
+rewrite -(cat_take_drop i b) -map_rev rev_cat !map_cat cat_path.
+case/andP=> ordb _; rewrite count_cat -{1}(size_takel (ltnW ltib)) ltnNge.
+rewrite addnC ((count _ _ =P 0) _) ?count_size //.
+rewrite eqn0Ngt -has_count; apply/hasPn=> y b_y /=; rewrite -leqNgt.
+apply: leq_trans le_x_e; have ->: x = last x (rev (drop i b)).
+ by rewrite (drop_nth 1 ltib) rev_cons last_rcons.
+rewrite -mem_rev in b_y; case/splitPr: (rev _) / b_y ordb => b1 b2.
+rewrite !map_cat cat_path last_cat /=; case/and3P=> _ _.
+move/(order_path_min leq_trans); case/lastP: b2 => // b3 x'.
+by move/allP; apply; rewrite ?map_f ?last_rcons ?mem_rcons ?mem_head.
+Qed.
+
+Lemma eq_abelian_type_isog G H :
+ abelian G -> abelian H -> isog G H = (abelian_type G == abelian_type H).
+Proof.
+move=> cGG cHH; apply/idP/eqP; first exact: isog_abelian_type.
+have{cGG} [bG defG <-] := abelian_structure cGG.
+have{cHH} [bH defH <-] := abelian_structure cHH.
+elim: bG bH G H defG defH => [|x bG IHb] [|y bH] // G H.
+ rewrite !big_nil => <- <- _.
+ by rewrite isog_cyclic_card ?cyclic1 ?cards1.
+rewrite !big_cons => defG defH /= [eqxy eqb].
+apply: (isog_dprod defG defH).
+ by rewrite isog_cyclic_card ?cycle_cyclic -?orderE ?eqxy /=.
+case/dprodP: defG => [[_ G' _ defG]] _ _ _; rewrite defG.
+case/dprodP: defH => [[_ H' _ defH]] _ _ _; rewrite defH.
+exact: IHb eqb.
+Qed.
+
+Lemma isog_abelem_card p G H :
+ p.-abelem G -> isog G H = p.-abelem H && (#|H| == #|G|).
+Proof.
+move=> abelG; apply/idP/andP=> [isoGH | [abelH eqGH]].
+ by rewrite -(isog_abelem isoGH) (card_isog isoGH).
+rewrite eq_abelian_type_isog ?(@abelem_abelian _ p) //.
+by rewrite !(@abelian_type_abelem _ p) ?(@rank_abelem _ p) // (eqP eqGH).
+Qed.
+
+Variables (D : {group aT}) (f : {morphism D >-> rT}).
+
+Lemma morphim_rank_abelian G : abelian G -> 'r(f @* G) <= 'r(G).
+Proof.
+move=> cGG; have sHG := subsetIr D G; apply: leq_trans (rankS sHG).
+rewrite -!grank_abelian ?morphim_abelian ?(abelianS sHG) //=.
+by rewrite -morphimIdom morphim_grank ?subsetIl.
+Qed.
+
+Lemma morphim_p_rank_abelian p G : abelian G -> 'r_p(f @* G) <= 'r_p(G).
+Proof.
+move=> cGG; have sHG := subsetIr D G; apply: leq_trans (p_rankS p sHG).
+have cHH := abelianS sHG cGG; rewrite -morphimIdom /=; set H := D :&: G.
+have sylP := nilpotent_pcore_Hall p (abelian_nil cHH).
+have sPH := pHall_sub sylP.
+have sPD: 'O_p(H) \subset D by rewrite (subset_trans sPH) ?subsetIl.
+rewrite -(p_rank_Sylow (morphim_pHall f sPD sylP)) -(p_rank_Sylow sylP) //.
+rewrite -!rank_pgroup ?morphim_pgroup ?pcore_pgroup //.
+by rewrite morphim_rank_abelian ?(abelianS sPH).
+Qed.
+
+Lemma isog_homocyclic G H : G \isog H -> homocyclic G = homocyclic H.
+Proof.
+move=> isoGH.
+by rewrite /homocyclic (isog_abelian isoGH) (isog_abelian_type isoGH).
+Qed.
+
+End IsogAbelian.
+
+Section QuotientRank.
+
+Variables (gT : finGroupType) (p : nat) (G H : {group gT}).
+Hypothesis cGG : abelian G.
+
+Lemma quotient_rank_abelian : 'r(G / H) <= 'r(G).
+Proof. exact: morphim_rank_abelian. Qed.
+
+Lemma quotient_p_rank_abelian : 'r_p(G / H) <= 'r_p(G).
+Proof. exact: morphim_p_rank_abelian. Qed.
+
+End QuotientRank.
+
+
+
+
diff --git a/mathcomp/solvable/all.v b/mathcomp/solvable/all.v
new file mode 100644
index 0000000..72405c5
--- /dev/null
+++ b/mathcomp/solvable/all.v
@@ -0,0 +1,19 @@
+Require Export abelian.
+Require Export alt.
+Require Export burnside_app.
+Require Export center.
+Require Export commutator.
+Require Export extraspecial.
+Require Export extremal.
+Require Export finmodule.
+Require Export frobenius.
+Require Export gfunctor.
+Require Export gseries.
+Require Export hall.
+Require Export jordanholder.
+Require Export maximal.
+Require Export nilpotent.
+Require Export pgroup.
+Require Export primitive_action.
+Require Export sylow.
+Require Export wielandt_fixpoint.
diff --git a/mathcomp/solvable/alt.v b/mathcomp/solvable/alt.v
new file mode 100644
index 0000000..3f12ad7
--- /dev/null
+++ b/mathcomp/solvable/alt.v
@@ -0,0 +1,528 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype tuple.
+Require Import tuple bigop prime finset fingroup morphism perm automorphism.
+Require Import quotient action cyclic pgroup gseries sylow primitive_action.
+
+(******************************************************************************)
+(* Definitions of the symmetric and alternate groups, and some properties. *)
+(* 'Sym_T == The symmetric group over type T (which must have a finType *)
+(* structure). *)
+(* := [set: {perm T}] *)
+(* 'Alt_T == The alternating group over type T. *)
+(******************************************************************************)
+
+Unset Printing Implicit Defensive.
+Set Implicit Arguments.
+Unset Strict Implicit.
+
+Import GroupScope.
+
+Definition bool_groupMixin := FinGroup.Mixin addbA addFb addbb.
+Canonical bool_baseGroup := Eval hnf in BaseFinGroupType bool bool_groupMixin.
+Canonical boolGroup := Eval hnf in FinGroupType addbb.
+
+Section SymAltDef.
+
+Variable T : finType.
+Implicit Types (s : {perm T}) (x y z : T).
+
+(** Definitions of the alternate groups and some Properties **)
+Definition Sym of phant T : {set {perm T}} := setT.
+
+Canonical Sym_group phT := Eval hnf in [group of Sym phT].
+
+Notation Local "'Sym_T" := (Sym (Phant T)) (at level 0).
+
+Canonical sign_morph := @Morphism _ _ 'Sym_T _ (in2W (@odd_permM _)).
+
+Definition Alt of phant T := 'ker (@odd_perm T).
+
+Canonical Alt_group phT := Eval hnf in [group of Alt phT].
+
+Notation Local "'Alt_T" := (Alt (Phant T)) (at level 0).
+
+Lemma Alt_even p : (p \in 'Alt_T) = ~~ p.
+Proof. by rewrite !inE /=; case: odd_perm. Qed.
+
+Lemma Alt_subset : 'Alt_T \subset 'Sym_T.
+Proof. exact: subsetT. Qed.
+
+Lemma Alt_normal : 'Alt_T <| 'Sym_T.
+Proof. exact: ker_normal. Qed.
+
+Lemma Alt_norm : 'Sym_T \subset 'N('Alt_T).
+Proof. by case/andP: Alt_normal. Qed.
+
+Let n := #|T|.
+
+Lemma Alt_index : 1 < n -> #|'Sym_T : 'Alt_T| = 2.
+Proof.
+move=> lt1n; rewrite -card_quotient ?Alt_norm //=.
+have : ('Sym_T / 'Alt_T) \isog (@odd_perm T @* 'Sym_T) by apply: first_isog.
+case/isogP=> g /injmP/card_in_imset <-.
+rewrite /morphim setIid=> ->; rewrite -card_bool; apply: eq_card => b.
+apply/imsetP; case: b => /=; last first.
+ by exists (1 : {perm T}); [rewrite setIid inE | rewrite odd_perm1].
+case: (pickP T) lt1n => [x1 _ | d0]; last by rewrite /n eq_card0.
+rewrite /n (cardD1 x1) ltnS lt0n => /existsP[x2 /=].
+by rewrite eq_sym andbT -odd_tperm; exists (tperm x1 x2); rewrite ?inE.
+Qed.
+
+Lemma card_Sym : #|'Sym_T| = n`!.
+Proof.
+rewrite -[n]cardsE -card_perm; apply: eq_card => p.
+by apply/idP/subsetP=> [? ?|]; rewrite !inE.
+Qed.
+
+Lemma card_Alt : 1 < n -> (2 * #|'Alt_T|)%N = n`!.
+Proof.
+by move/Alt_index <-; rewrite mulnC (Lagrange Alt_subset) card_Sym.
+Qed.
+
+Lemma Sym_trans : [transitive^n 'Sym_T, on setT | 'P].
+Proof.
+apply/imsetP; pose t1 := [tuple of enum T].
+have dt1: t1 \in n.-dtuple(setT) by rewrite inE enum_uniq; apply/subsetP.
+exists t1 => //; apply/setP=> t; apply/idP/imsetP=> [|[a _ ->{t}]]; last first.
+ by apply: n_act_dtuple => //; apply/astabsP=> x; rewrite !inE.
+case/dtuple_onP=> injt _; have injf := inj_comp injt enum_rank_inj.
+exists (perm injf); first by rewrite inE.
+apply: eq_from_tnth => i; rewrite tnth_map /= [aperm _ _]permE; congr tnth.
+by rewrite (tnth_nth (enum_default i)) enum_valK.
+Qed.
+
+Lemma Alt_trans : [transitive^n.-2 'Alt_T, on setT | 'P].
+Proof.
+case n_m2: n Sym_trans => [|[|m]] /= tr_m2; try exact: ntransitive0.
+have tr_m := ntransitive_weak (leqW (leqnSn m)) tr_m2.
+case/imsetP: tr_m2; case/tupleP=> x; case/tupleP=> y t.
+rewrite !dtuple_on_add 2![x \in _]inE inE negb_or /= -!andbA.
+case/and4P=> nxy ntx nty dt _; apply/imsetP; exists t => //; apply/setP=> u.
+apply/idP/imsetP=> [|[a _ ->{u}]]; last first.
+ by apply: n_act_dtuple => //; apply/astabsP=> z; rewrite !inE.
+case/(atransP2 tr_m dt)=> /= a _ ->{u}.
+case odd_a: (odd_perm a); last by exists a => //; rewrite !inE /= odd_a.
+exists (tperm x y * a); first by rewrite !inE /= odd_permM odd_tperm nxy odd_a.
+apply/val_inj/eq_in_map => z tz; rewrite actM /= /aperm; congr (a _).
+by case: tpermP ntx nty => // <-; rewrite tz.
+Qed.
+
+Lemma aperm_faithful (A : {group {perm T}}) : [faithful A, on setT | 'P].
+Proof.
+by apply/faithfulP=> /= p _ np1; apply/eqP/perm_act1P=> y; rewrite np1 ?inE.
+Qed.
+
+End SymAltDef.
+
+Notation "''Sym_' T" := (Sym (Phant T))
+ (at level 8, T at level 2, format "''Sym_' T") : group_scope.
+Notation "''Sym_' T" := (Sym_group (Phant T)) : Group_scope.
+
+Notation "''Alt_' T" := (Alt (Phant T))
+ (at level 8, T at level 2, format "''Alt_' T") : group_scope.
+Notation "''Alt_' T" := (Alt_group (Phant T)) : Group_scope.
+
+Lemma trivial_Alt_2 (T : finType) : #|T| <= 2 -> 'Alt_T = 1.
+Proof.
+rewrite leq_eqVlt => /predU1P[] oT.
+ by apply: card_le1_trivg; rewrite -leq_double -mul2n card_Alt oT.
+suffices Sym1: 'Sym_T = 1 by apply/trivgP; rewrite -Sym1 subsetT.
+by apply: card1_trivg; rewrite card_Sym; case: #|T| oT; do 2?case.
+Qed.
+
+Lemma simple_Alt_3 (T : finType) : #|T| = 3 -> simple 'Alt_T.
+Proof.
+move=> T3; have{T3} oA: #|'Alt_T| = 3.
+ by apply: double_inj; rewrite -mul2n card_Alt T3.
+apply/simpleP; split=> [|K]; [by rewrite trivg_card1 oA | case/andP=> sKH _].
+have:= cardSg sKH; rewrite oA dvdn_divisors // !inE orbC /= -oA.
+case/pred2P=> eqK; [right | left]; apply/eqP.
+ by rewrite eqEcard sKH eqK leqnn.
+by rewrite eq_sym eqEcard sub1G eqK cards1.
+Qed.
+
+Lemma not_simple_Alt_4 (T : finType) : #|T| = 4 -> ~~ simple 'Alt_T.
+Proof.
+move=> oT; set A := 'Alt_T.
+have oA: #|A| = 12 by apply: double_inj; rewrite -mul2n card_Alt oT.
+suffices [p]: exists p, [/\ prime p, 1 < #|A|`_p < #|A| & #|'Syl_p(A)| == 1%N].
+ case=> p_pr pA_int; rewrite /A; case/normal_sylowP=> P; case/pHallP.
+ rewrite /= -/A => sPA pP nPA; apply/simpleP=> [] [_]; rewrite -pP in pA_int.
+ by case/(_ P)=> // defP; rewrite defP oA ?cards1 in pA_int.
+have: #|'Syl_3(A)| \in filter [pred d | d %% 3 == 1%N] (divisors 12).
+ by rewrite mem_filter -dvdn_divisors //= -oA card_Syl_dvd ?card_Syl_mod.
+rewrite /= oA mem_seq2 orbC.
+case/predU1P=> [oQ3|]; [exists 2 | exists 3]; split; rewrite ?p_part //.
+pose A3 := [set x : {perm T} | #[x] == 3]; suffices oA3: #|A :&: A3| = 8.
+ have sQ2 P: P \in 'Syl_2(A) -> P :=: A :\: A3.
+ rewrite inE pHallE oA p_part -natTrecE /= => /andP[sPA /eqP oP].
+ apply/eqP; rewrite eqEcard -(leq_add2l 8) -{1}oA3 cardsID oA oP.
+ rewrite andbT subsetD sPA; apply/exists_inP=> -[x] /= Px.
+ by rewrite inE => /eqP ox; have:= order_dvdG Px; rewrite oP ox.
+ have [/= P sylP] := Sylow_exists 2 [group of A].
+ rewrite -(([set P] =P 'Syl_2(A)) _) ?cards1 // eqEsubset sub1set inE sylP.
+ by apply/subsetP=> Q sylQ; rewrite inE -val_eqE /= !sQ2 // inE.
+rewrite -[8]/(4 * 2)%N -{}oQ3 -sum1_card -sum_nat_const.
+rewrite (partition_big (fun x => <[x]>%G) (mem 'Syl_3(A))) => [|x]; last first.
+ by case/setIP=> Ax; rewrite /= !inE pHallE p_part cycle_subG Ax oA.
+apply: eq_bigr => Q; rewrite inE /= inE pHallE oA p_part -?natTrecE //=.
+case/andP=> sQA /eqP oQ; have:= oQ.
+rewrite (cardsD1 1) group1 -sum1_card => [[/= <-]]; apply: eq_bigl => x.
+rewrite setIC -val_eqE /= 2!inE in_setD1 -andbA -{4}[x]expg1 -order_dvdn dvdn1.
+apply/and3P/andP=> [[/eqP-> _ /eqP <-] | [ntx Qx]]; first by rewrite cycle_id.
+have:= order_dvdG Qx; rewrite oQ dvdn_divisors // mem_seq2 (negPf ntx) /=.
+by rewrite eqEcard cycle_subG Qx (subsetP sQA) // oQ /order => /eqP->.
+Qed.
+
+Lemma simple_Alt5_base (T : finType) : #|T| = 5 -> simple 'Alt_T.
+Proof.
+move=> oT.
+have F1: #|'Alt_T| = 60 by apply: double_inj; rewrite -mul2n card_Alt oT.
+have FF (H : {group {perm T}}): H <| 'Alt_T -> H :<>: 1 -> 20 %| #|H|.
+- move=> Hh1 Hh3.
+ have [x _]: exists x, x \in T by apply/existsP/eqP; rewrite oT.
+ have F2 := Alt_trans T; rewrite oT /= in F2.
+ have F3: [transitive 'Alt_T, on setT | 'P] by exact: ntransitive1 F2.
+ have F4: [primitive 'Alt_T, on setT | 'P] by exact: ntransitive_primitive F2.
+ case: (prim_trans_norm F4 Hh1) => F5.
+ case: Hh3; apply/trivgP; exact: subset_trans F5 (aperm_faithful _).
+ have F6: 5 %| #|H| by rewrite -oT -cardsT (atrans_dvd F5).
+ have F7: 4 %| #|H|.
+ have F7: #|[set~ x]| = 4 by rewrite cardsC1 oT.
+ case: (pickP (mem [set~ x])) => [y Hy | ?]; last by rewrite eq_card0 in F7.
+ pose K := 'C_H[x | 'P]%G.
+ have F8 : K \subset H by apply: subsetIl.
+ pose Gx := 'C_('Alt_T)[x | 'P]%G.
+ have F9: [transitive^2 Gx, on [set~ x] | 'P].
+ by rewrite -[[set~ x]]setTI -setDE stab_ntransitive ?inE.
+ have F10: [transitive Gx, on [set~ x] | 'P].
+ exact: ntransitive1 F9.
+ have F11: [primitive Gx, on [set~ x] | 'P].
+ exact: ntransitive_primitive F9.
+ have F12: K \subset Gx by apply: setSI; exact: normal_sub.
+ have F13: K <| Gx by rewrite /(K <| _) F12 normsIG // normal_norm.
+ case: (prim_trans_norm F11 F13) => Ksub; last first.
+ apply: dvdn_trans (cardSg F8); rewrite -F7; exact: atrans_dvd Ksub.
+ have F14: [faithful Gx, on [set~ x] | 'P].
+ apply/subsetP=> g; do 2![case/setIP] => Altg cgx cgx'.
+ apply: (subsetP (aperm_faithful 'Alt_T)).
+ rewrite inE Altg /=; apply/astabP=> z _.
+ case: (z =P x) => [->|]; first exact: (astab1P cgx).
+ by move/eqP=> nxz; rewrite (astabP cgx') ?inE //.
+ have Hreg g (z : T): g \in H -> g z = z -> g = 1.
+ have F15 h: h \in H -> h x = x -> h = 1.
+ move=> Hh Hhx; have: h \in K by rewrite inE Hh; apply/astab1P.
+ by rewrite (trivGP (subset_trans Ksub F14)) => /set1P.
+ move=> Hg Hgz; have:= in_setT x; rewrite -(atransP F3 z) ?inE //.
+ case/imsetP=> g1 Hg1 Hg2; apply: (conjg_inj g1); rewrite conj1g.
+ apply: F15; last by rewrite Hg2 -permM mulKVg permM Hgz.
+ by case/normalP: Hh1 => _ nH1; rewrite -(nH1 _ Hg1) memJ_conjg.
+ clear K F8 F12 F13 Ksub F14.
+ case: (Cauchy _ F6) => // h Hh /eqP Horder.
+ have diff_hnx_x n: 0 < n -> n < 5 -> x != (h ^+ n) x.
+ move=> Hn1 Hn2; rewrite eq_sym; apply/negP => HH.
+ have: #[h ^+ n] = 5.
+ rewrite orderXgcd // (eqP Horder).
+ by move: Hn1 Hn2 {HH}; do 5 (case: n => [|n] //).
+ have Hhd2: h ^+ n \in H by rewrite groupX.
+ by rewrite (Hreg _ _ Hhd2 (eqP HH)) order1.
+ pose S1 := [tuple x; h x; (h ^+ 3) x].
+ have DnS1: S1 \in 3.-dtuple(setT).
+ rewrite inE memtE subset_all /= !inE /= !negb_or -!andbA /= andbT.
+ rewrite -{1}[h]expg1 !diff_hnx_x // expgSr permM.
+ by rewrite (inj_eq (@perm_inj _ _)) diff_hnx_x.
+ pose S2 := [tuple x; h x; (h ^+ 2) x].
+ have DnS2: S2 \in 3.-dtuple(setT).
+ rewrite inE memtE subset_all /= !inE /= !negb_or -!andbA /= andbT.
+ rewrite -{1}[h]expg1 !diff_hnx_x // expgSr permM.
+ by rewrite (inj_eq (@perm_inj _ _)) diff_hnx_x.
+ case: (atransP2 F2 DnS1 DnS2) => g Hg [/=].
+ rewrite /aperm => Hgx Hghx Hgh3x.
+ have h_g_com: h * g = g * h.
+ suff HH: (g * h * g^-1) * h^-1 = 1 by rewrite -[h * g]mul1g -HH !gnorm.
+ apply: (Hreg _ x); last first.
+ by rewrite !permM -Hgx Hghx -!permM mulKVg mulgV perm1.
+ rewrite groupM // ?groupV // (conjgCV g) mulgK -mem_conjg.
+ by case/normalP: Hh1 => _ ->.
+ have: (g * (h ^+ 2) * g ^-1) x = (h ^+ 3) x.
+ rewrite !permM -Hgx.
+ have ->: h (h x) = (h ^+ 2) x by rewrite /= permM.
+ by rewrite {1}Hgh3x -!permM /= mulgV mulg1 -expgSr.
+ rewrite commuteX // mulgK {1}[expgn]lock expgS permM -lock.
+ by move/perm_inj=> eqxhx; case/eqP: (diff_hnx_x 1%N isT isT); rewrite expg1.
+ by rewrite (@Gauss_dvd 4 5) // F7.
+apply/simpleP; split => [|H Hnorm]; first by rewrite trivg_card1 F1.
+case Hcard1: (#|H| == 1%N); move/eqP: Hcard1 => Hcard1.
+ by left; apply: card1_trivg; rewrite Hcard1.
+right; case Hcard60: (#|H| == 60%N); move/eqP: Hcard60 => Hcard60.
+ by apply/eqP; rewrite eqEcard Hcard60 F1 andbT; case/andP: Hnorm.
+have Hcard20: #|H| = 20; last clear Hcard1 Hcard60.
+ have Hdiv: 20 %| #|H| by apply: FF => // HH; case Hcard1; rewrite HH cards1.
+ case H20: (#|H| == 20); first by apply/eqP.
+ case: Hcard60; case/andP: Hnorm; move/cardSg; rewrite F1 => Hdiv1 _.
+ by case/dvdnP: Hdiv H20 Hdiv1 => n ->; move: n; do 4!case=> //.
+have prime_5: prime 5 by [].
+have nSyl5: #|'Syl_5(H)| = 1%N.
+ move: (card_Syl_dvd 5 H) (card_Syl_mod H prime_5).
+ rewrite Hcard20; case: (card _) => // n Hdiv.
+ move: (dvdn_leq (isT: (0 < 20)%N) Hdiv).
+ by move: (n) Hdiv; do 20 (case => //).
+case: (Sylow_exists 5 H) => S; case/pHallP=> sSH oS.
+have{oS} oS: #|S| = 5 by rewrite oS p_part Hcard20.
+suff: 20 %| #|S| by rewrite oS.
+apply FF => [|S1]; last by rewrite S1 cards1 in oS.
+apply: char_normal_trans Hnorm; apply: lone_subgroup_char => // Q sQH isoQS.
+rewrite subEproper; apply/norP=> [[nQS _]]; move: nSyl5.
+rewrite (cardsD1 S) (cardsD1 Q) 4!{1}inE nQS !pHallE sQH sSH Hcard20 p_part.
+by rewrite (card_isog isoQS) oS.
+Qed.
+
+Section Restrict.
+
+Variables (T : finType) (x : T).
+Notation T' := {y | y != x}.
+
+Lemma rfd_funP (p : {perm T}) (u : T') :
+ let p1 := if p x == x then p else 1 in p1 (val u) != x.
+Proof.
+case: (p x =P x) => /= [pxx|_]; last by rewrite perm1 (valP u).
+by rewrite -{2}pxx (inj_eq (@perm_inj _ p)); exact: (valP u).
+Qed.
+
+Definition rfd_fun p := [fun u => Sub ((_ : {perm T}) _) (rfd_funP p u) : T'].
+
+Lemma rfdP p : injective (rfd_fun p).
+Proof.
+apply: can_inj (rfd_fun p^-1) _ => u; apply: val_inj => /=.
+rewrite -(inj_eq (@perm_inj _ p)) permKV eq_sym.
+by case: eqP => _; rewrite !(perm1, permK).
+Qed.
+
+Definition rfd p := perm (@rfdP p).
+
+Hypothesis card_T : 2 < #|T|.
+
+Lemma rfd_morph : {in 'C_('Sym_T)[x | 'P] &, {morph rfd : y z / y * z}}.
+Proof.
+move=> p q; rewrite !setIA !setIid; move/astab1P=> p_x; move/astab1P=> q_x.
+apply/permP=> u; apply: val_inj.
+by rewrite permE /= !permM !permE /= [p x]p_x [q x]q_x eqxx permM /=.
+Qed.
+
+Canonical rfd_morphism := Morphism rfd_morph.
+
+Definition rgd_fun (p : {perm T'}) :=
+ [fun x1 => if insub x1 is Some u then sval (p u) else x].
+
+Lemma rgdP p : injective (rgd_fun p).
+Proof.
+apply: can_inj (rgd_fun p^-1) _ => y /=.
+case: (insubP _ y) => [u _ val_u|]; first by rewrite valK permK.
+by rewrite negbK; move/eqP->; rewrite insubF //= eqxx.
+Qed.
+
+Definition rgd p := perm (@rgdP p).
+
+Lemma rfd_odd (p : {perm T}) : p x = x -> rfd p = p :> bool.
+Proof.
+have rfd1: rfd 1 = 1.
+ by apply/permP => u; apply: val_inj; rewrite permE /= if_same !perm1.
+have HP0 (t : {perm T}): #|[set x | t x != x]| = 0 -> rfd t = t :> bool.
+- move=> Ht; suff ->: t = 1 by rewrite rfd1 !odd_perm1.
+ apply/permP => z; rewrite perm1; apply/eqP/wlog_neg => nonfix_z.
+ by rewrite (cardD1 z) inE nonfix_z in Ht.
+elim: #|_| {-2}p (leqnn #|[set x | p x != x]|) => {p}[|n Hrec] p Hp Hpx.
+ by apply: HP0; move: Hp; case: card.
+case Ex: (pred0b (mem [set x | p x != x])); first by apply: HP0; move/eqnP: Ex.
+case/pred0Pn: Ex => x1; rewrite /= inE => Hx1.
+have nx1x: x1 != x by apply/eqP => HH; rewrite HH Hpx eqxx in Hx1.
+have npxx1: p x != x1 by apply/eqP => HH; rewrite -HH !Hpx eqxx in Hx1.
+have npx1x: p x1 != x.
+ by apply/eqP; rewrite -Hpx; move/perm_inj => HH; case/eqP: nx1x.
+pose p1 := p * tperm x1 (p x1).
+have Hp1: p1 x = x.
+ by rewrite /p1 permM; case tpermP => // [<-|]; [rewrite Hpx | move/perm_inj].
+have Hcp1: #|[set x | p1 x != x]| <= n.
+ have F1 y: p y = y -> p1 y = y.
+ move=> Hy; rewrite /p1 permM Hy.
+ case tpermP => //; first by move => <-.
+ by move=> Hpx1; apply: (@perm_inj _ p); rewrite -Hpx1.
+ have F2: p1 x1 = x1 by rewrite /p1 permM tpermR.
+ have F3: [set x | p1 x != x] \subset [predD1 [set x | p x != x] & x1].
+ apply/subsetP => z; rewrite !inE permM.
+ case tpermP => HH1 HH2.
+ - rewrite eq_sym HH1 andbb; apply/eqP=> dx1.
+ by rewrite dx1 HH1 dx1 eqxx in HH2.
+ - by rewrite (perm_inj HH1) eqxx in HH2.
+ by move->; rewrite andbT; apply/eqP => HH3; rewrite HH3 in HH2.
+ apply: (leq_trans (subset_leq_card F3)).
+ by move: Hp; rewrite (cardD1 x1) inE Hx1.
+have ->: p = p1 * tperm x1 (p x1) by rewrite -mulgA tperm2 mulg1.
+rewrite odd_permM odd_tperm eq_sym Hx1 morphM; last 2 first.
+- by rewrite 2!inE; exact/astab1P.
+- by rewrite 2!inE; apply/astab1P; rewrite -{1}Hpx /= /aperm -permM.
+rewrite odd_permM Hrec //=; congr (_ (+) _).
+pose x2 : T' := Sub x1 nx1x; pose px2 : T' := Sub (p x1) npx1x.
+suff ->: rfd (tperm x1 (p x1)) = tperm x2 px2.
+ by rewrite odd_tperm -val_eqE eq_sym.
+apply/permP => z; apply/val_eqP; rewrite permE /= tpermD // eqxx.
+case: (tpermP x2) => [->|->|HH1 HH2]; rewrite /x2 ?tpermL ?tpermR 1?tpermD //.
+ by apply/eqP=> HH3; case: HH1; apply: val_inj.
+by apply/eqP => HH3; case: HH2; apply: val_inj.
+Qed.
+
+Lemma rfd_iso : 'C_('Alt_T)[x | 'P] \isog 'Alt_T'.
+Proof.
+have rgd_x p: rgd p x = x by rewrite permE /= insubF //= eqxx.
+have rfd_rgd p: rfd (rgd p) = p.
+ apply/permP => [[z Hz]]; apply/val_eqP; rewrite !permE.
+ rewrite /= [rgd _ _]permE /= insubF eq_refl // permE /=.
+ by rewrite (@insubT _ (xpredC1 x) _ _ Hz).
+have sSd: 'C_('Alt_T)[x | 'P] \subset 'dom rfd.
+ by apply/subsetP=> p; rewrite !inE /=; case/andP.
+apply/isogP; exists [morphism of restrm sSd rfd] => /=; last first.
+ rewrite morphim_restrm setIid; apply/setP=> z; apply/morphimP/idP=> [[p _]|].
+ case/setIP; rewrite Alt_even => Hp; move/astab1P=> Hp1 ->.
+ by rewrite Alt_even rfd_odd.
+ have dz': rgd z x == x by rewrite rgd_x.
+ move=> kz; exists (rgd z); last by rewrite /= rfd_rgd.
+ by rewrite 2!inE (sameP astab1P eqP).
+ rewrite 4!inE /= (sameP astab1P eqP) dz' -rfd_odd; last exact/eqP.
+ by rewrite rfd_rgd mker // ?set11.
+apply/injmP=> x1 y1 /=.
+case/setIP=> Hax1; move/astab1P; rewrite /= /aperm => Hx1.
+case/setIP=> Hay1; move/astab1P; rewrite /= /aperm => Hy1 Hr.
+apply/permP => z.
+case (z =P x) => [->|]; [by rewrite Hx1 | move/eqP => nzx].
+move: (congr1 (fun q : {perm T'} => q (Sub z nzx)) Hr).
+by rewrite !permE => [[]]; rewrite Hx1 Hy1 !eqxx.
+Qed.
+
+End Restrict.
+
+Lemma simple_Alt5 (T : finType) : #|T| >= 5 -> simple 'Alt_T.
+Proof.
+suff F1 n: #|T| = n + 5 -> simple 'Alt_T by move/subnK/esym/F1.
+elim: n T => [| n Hrec T Hde]; first exact: simple_Alt5_base.
+have oT: 5 < #|T| by rewrite Hde addnC.
+apply/simpleP; split=> [|H Hnorm]; last have [Hh1 nH] := andP Hnorm.
+ rewrite trivg_card1 -[#|_|]half_double -mul2n card_Alt Hde addnC //.
+ by rewrite addSn factS mulnC -(prednK (fact_gt0 _)).
+case E1: (pred0b T); first by rewrite /pred0b in E1; rewrite (eqP E1) in oT.
+case/pred0Pn: E1 => x _; have Hx := in_setT x.
+have F2: [transitive^4 'Alt_T, on setT | 'P].
+ by apply: ntransitive_weak (Alt_trans T); rewrite -(subnKC oT).
+have F3 := ntransitive1 (isT: 0 < 4) F2.
+have F4 := ntransitive_primitive (isT: 1 < 4) F2.
+case Hcard1: (#|H| == 1%N); move/eqP: Hcard1 => Hcard1.
+ by left; apply: card1_trivg; rewrite Hcard1.
+right; case: (prim_trans_norm F4 Hnorm) => F5.
+ by rewrite (trivGP (subset_trans F5 (aperm_faithful _))) cards1 in Hcard1.
+case E1: (pred0b (predD1 T x)).
+ rewrite /pred0b in E1; move: oT.
+ by rewrite (cardD1 x) (eqP E1); case: (T x).
+case/pred0Pn: E1 => y Hdy; case/andP: (Hdy) => diff_x_y Hy.
+pose K := 'C_H[x | 'P]%G.
+have F8: K \subset H by apply: subsetIl.
+pose Gx := 'C_('Alt_T)[x | 'P].
+have F9: [transitive^3 Gx, on [set~ x] | 'P].
+ by rewrite -[[set~ x]]setTI -setDE stab_ntransitive ?inE.
+have F10: [transitive Gx, on [set~ x] | 'P].
+ by apply: ntransitive1 F9.
+have F11: [primitive Gx, on [set~ x] | 'P].
+ by apply: ntransitive_primitive F9.
+have F12: K \subset Gx by rewrite setSI // normal_sub.
+have F13: K <| Gx by apply/andP; rewrite normsIG.
+have:= prim_trans_norm F11; case/(_ K) => //= => Ksub; last first.
+ have F14: Gx * H = 'Alt_T by exact/(subgroup_transitiveP _ _ F3).
+ have: simple Gx.
+ by rewrite (isog_simple (rfd_iso x)) Hrec //= card_sig cardC1 Hde.
+ case/simpleP=> _ simGx; case/simGx: F13 => /= HH2.
+ case Ez: (pred0b (predD1 (predD1 T x) y)).
+ move: oT; rewrite /pred0b in Ez.
+ by rewrite (cardD1 x) (cardD1 y) (eqP Ez) inE /= inE /= diff_x_y.
+ case/pred0Pn: Ez => z; case/andP => diff_y_z Hdz.
+ have [diff_x_z Hz] := andP Hdz.
+ have: z \in [set~ x] by rewrite !inE.
+ rewrite -(atransP Ksub y) ?inE //; case/imsetP => g.
+ rewrite /= HH2 inE; move/eqP=> -> HH4.
+ by case/negP: diff_y_z; rewrite HH4 act1.
+ by rewrite /= -F14 -[Gx]HH2 (mulSGid F8).
+have F14: [faithful Gx, on [set~ x] | 'P].
+ apply: subset_trans (aperm_faithful 'Sym_T); rewrite subsetI subsetT.
+ apply/subsetP=> g; do 2![case/setIP]=> _ cgx cgx'; apply/astabP=> z _ /=.
+ case: (z =P x) => [->|]; first exact: (astab1P cgx).
+ by move/eqP=> zx; rewrite [_ g](astabP cgx') ?inE.
+have Hreg g z: g \in H -> g z = z -> g = 1.
+ have F15 h: h \in H -> h x = x -> h = 1.
+ move=> Hh Hhx; have: h \in K by rewrite inE Hh; apply/astab1P.
+ by rewrite [K](trivGP (subset_trans Ksub F14)) => /set1P.
+ move=> Hg Hgz; have:= in_setT x; rewrite -(atransP F3 z) ?inE //.
+ case/imsetP=> g1 Hg1 Hg2; apply: (conjg_inj g1); rewrite conj1g.
+ apply: F15; last by rewrite Hg2 -permM mulKVg permM Hgz.
+ by rewrite memJ_norm ?(subsetP nH).
+clear K F8 F12 F13 Ksub F14.
+have Hcard: 5 < #|H|.
+ apply: (leq_trans oT); apply dvdn_leq; first by exact: cardG_gt0.
+ by rewrite -cardsT (atrans_dvd F5).
+case Eh: (pred0b [predD1 H & 1]).
+ by move: Hcard; rewrite /pred0b in Eh; rewrite (cardD1 1) group1 (eqP Eh).
+case/pred0Pn: Eh => h; case/andP => diff_1_h /= Hh.
+case Eg: (pred0b (predD1 (predD1 [predD1 H & 1] h) h^-1)).
+ move: Hcard; rewrite ltnNge; case/negP.
+ rewrite (cardD1 1) group1 (cardD1 h) (cardD1 h^-1) (eqnP Eg).
+ by do 2!case: (_ \in _).
+case/pred0Pn: Eg => g; case/andP => diff_h1_g; case/andP => diff_h_g.
+case/andP => diff_1_g /= Hg.
+case diff_hx_x: (h x == x).
+by case/negP: diff_1_h; apply/eqP; apply: (Hreg _ _ Hh (eqP diff_hx_x)).
+case diff_gx_x: (g x == x).
+ case/negP: diff_1_g; apply/eqP; apply: (Hreg _ _ Hg (eqP diff_gx_x)).
+case diff_gx_hx: (g x == h x).
+ case/negP: diff_h_g; apply/eqP; symmetry; apply: (mulIg g^-1); rewrite gsimp.
+ apply: (Hreg _ x); first by rewrite groupM // groupV.
+ by rewrite permM -(eqP diff_gx_hx) -permM mulgV perm1.
+case diff_hgx_x: ((h * g) x == x).
+ case/negP: diff_h1_g; apply/eqP; apply: (mulgI h); rewrite !gsimp.
+ by apply: (Hreg _ x); [exact: groupM | apply/eqP].
+case diff_hgx_hx: ((h * g) x == h x).
+ case/negP: diff_1_g; apply/eqP.
+ by apply: (Hreg _ (h x)) => //; apply/eqP; rewrite -permM.
+case diff_hgx_gx: ((h * g) x == g x).
+ case/eqP: diff_hx_x; apply: (@perm_inj _ g) => //.
+ by apply/eqP; rewrite -permM.
+case Ez: (pred0b
+ (predD1 (predD1 (predD1 (predD1 T x) (h x)) (g x)) ((h * g) x))).
+- move: oT; rewrite /pred0b in Ez.
+ rewrite (cardD1 x) (cardD1 (h x)) (cardD1 (g x)) (cardD1 ((h * g) x)).
+ by rewrite (eqP Ez) addnC; do 3!case: (_ x \in _).
+case/pred0Pn: Ez => z.
+case/and5P=> diff_hgx_z diff_gx_z diff_hx_z diff_x_z /= Hz.
+pose S1 := [tuple x; h x; g x; z].
+have DnS1: S1 \in 4.-dtuple(setT).
+ rewrite inE memtE subset_all -!andbA !negb_or /= !inE !andbT.
+ rewrite -!(eq_sym z) diff_gx_z diff_x_z diff_hx_z.
+ by rewrite !(eq_sym x) diff_hx_x diff_gx_x eq_sym diff_gx_hx.
+pose S2 := [tuple x; h x; g x; (h * g) x].
+have DnS2: S2 \in 4.-dtuple(setT).
+ rewrite inE memtE subset_all -!andbA !negb_or /= !inE !andbT !(eq_sym x).
+ rewrite diff_hx_x diff_gx_x diff_hgx_x.
+ by rewrite !(eq_sym (h x)) diff_gx_hx diff_hgx_hx eq_sym diff_hgx_gx.
+case: (atransP2 F2 DnS1 DnS2) => k Hk [/=].
+rewrite /aperm => Hkx Hkhx Hkgx Hkhgx.
+have h_k_com: h * k = k * h.
+ suff HH: (k * h * k^-1) * h^-1 = 1 by rewrite -[h * k]mul1g -HH !gnorm.
+ apply: (Hreg _ x); last first.
+ by rewrite !permM -Hkx Hkhx -!permM mulKVg mulgV perm1.
+ by rewrite groupM // ?groupV // (conjgCV k) mulgK -mem_conjg (normsP nH).
+have g_k_com: g * k = k * g.
+ suff HH: (k * g * k^-1) * g^-1 = 1 by rewrite -[g * k]mul1g -HH !gnorm.
+ apply: (Hreg _ x); last first.
+ by rewrite !permM -Hkx Hkgx -!permM mulKVg mulgV perm1.
+ by rewrite groupM // ?groupV // (conjgCV k) mulgK -mem_conjg (normsP nH).
+have HH: (k * (h * g) * k ^-1) x = z.
+ by rewrite 2!permM -Hkx Hkhgx -permM mulgV perm1.
+case/negP: diff_hgx_z.
+rewrite -HH !mulgA -h_k_com -!mulgA [k * _]mulgA.
+by rewrite -g_k_com -!mulgA mulgV mulg1.
+Qed.
diff --git a/mathcomp/solvable/burnside_app.v b/mathcomp/solvable/burnside_app.v
new file mode 100644
index 0000000..8fe0b3f
--- /dev/null
+++ b/mathcomp/solvable/burnside_app.v
@@ -0,0 +1,1305 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import tuple finfun bigop finset fingroup action perm primitive_action.
+
+(* Application of the Burside formula to count the number of distinct *)
+(* colorings of the vertices of a square and a cube. *)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Lemma burnside_formula : forall (gT : finGroupType) s (G : {group gT}),
+ uniq s -> s =i G ->
+ forall (sT : finType) (to : {action gT &-> sT}),
+ (#|orbit to G @: setT| * size s)%N = \sum_(p <- s) #|'Fix_to[p]|.
+Proof.
+move=> gT s G Us sG sT to.
+rewrite big_uniq // -(card_uniqP Us) (eq_card sG) -Frobenius_Cauchy.
+ by apply: eq_big => // p _; rewrite setTI.
+by apply/actsP=> ? _ ?; rewrite !inE.
+Qed.
+
+Implicit Arguments burnside_formula [gT].
+
+Section colouring.
+
+Variable n : nat.
+Definition colors := 'I_n.
+Canonical colors_eqType := Eval hnf in [eqType of colors].
+Canonical colors_choiceType := Eval hnf in [choiceType of colors].
+Canonical colors_countType := Eval hnf in [countType of colors].
+Canonical colors_finType := Eval hnf in [finType of colors].
+
+Section square_colouring.
+
+Definition square := 'I_4.
+Canonical square_eqType := Eval hnf in [eqType of square].
+Canonical square_choiceType := Eval hnf in [choiceType of square].
+Canonical square_countType := Eval hnf in [countType of square].
+Canonical square_finType := Eval hnf in [finType of square].
+Canonical square_subType := Eval hnf in [subType of square].
+Canonical square_subCountType :=
+ Eval hnf in [subCountType of square].
+Canonical square_subFinType := Eval hnf in [subFinType of square].
+
+Definition mksquare i : square := Sub (i %% _) (ltn_mod i 4).
+Definition c0 := mksquare 0.
+Definition c1 := mksquare 1.
+Definition c2 := mksquare 2.
+Definition c3 := mksquare 3.
+
+(*rotations*)
+Definition R1 (sc : square) : square := tnth [tuple c1; c2; c3; c0] sc.
+
+Definition R2 (sc : square) : square := tnth [tuple c2; c3; c0; c1] sc.
+
+Definition R3 (sc : square) : square := tnth [tuple c3; c0; c1; c2] sc.
+
+Ltac get_inv elt l :=
+ match l with
+ | (_, (elt, ?x)) => x
+ | (elt, ?x) => x
+ | (?x, _) => get_inv elt x
+ end.
+
+Definition rot_inv := ((R1, R3), (R2, R2), (R3, R1)).
+
+Ltac inj_tac :=
+ move: (erefl rot_inv); unfold rot_inv;
+ match goal with |- ?X = _ -> injective ?Y =>
+ move=> _; let x := get_inv Y X in
+ apply: (can_inj (g:=x)); move => [val H1]
+ end.
+
+Lemma R1_inj : injective R1.
+Proof. by inj_tac; repeat (destruct val => //=; first by apply /eqP). Qed.
+
+Lemma R2_inj : injective R2.
+Proof. by inj_tac; repeat (destruct val => //=; first by apply /eqP). Qed.
+
+Lemma R3_inj : injective R3.
+Proof. by inj_tac; repeat (destruct val => //=; first by apply /eqP). Qed.
+
+Definition r1 := (perm R1_inj).
+Definition r2 := (perm R2_inj).
+Definition r3 := (perm R3_inj).
+Definition id1 := (1 : {perm square}).
+
+Definition is_rot (r : {perm _}) := (r * r1 == r1 * r).
+Definition rot := [set r | is_rot r].
+
+Lemma group_set_rot : group_set rot.
+Proof.
+apply /group_setP;split; first by rewrite /rot inE /is_rot mulg1 mul1g.
+move => x1 y; rewrite /rot !inE /= /is_rot; move/eqP => hx1; move/eqP => hy.
+by rewrite -mulgA hy !mulgA hx1.
+Qed.
+
+Canonical rot_group := Group group_set_rot.
+
+Definition rotations := [set id1; r1; r2; r3].
+
+Lemma rot_eq_c0 : forall r s : {perm square},
+ is_rot r -> is_rot s -> r c0 = s c0 -> r = s.
+Proof.
+rewrite /is_rot => r s; move/eqP => hr; move/eqP=> hs hrs; apply/permP => a.
+have ->: a = (r1 ^+ a) c0
+ by apply/eqP; case: a; do 4?case => //=; rewrite ?permM !permE.
+by rewrite -!permM -!commuteX // !permM hrs.
+Qed.
+
+Lemma rot_r1 : forall r, is_rot r -> r = r1 ^+ (r c0).
+Proof.
+move=> r hr;apply: rot_eq_c0 => //;apply/eqP.
+ by symmetry; exact: commuteX.
+by case: (r c0); do 4?case => //=; rewrite ?permM !permE /=.
+Qed.
+
+Lemma rotations_is_rot : forall r, r \in rotations -> is_rot r.
+Proof.
+move=> r Dr; apply/eqP; apply/permP => a; rewrite !inE -!orbA !permM in Dr *.
+by case/or4P: Dr; move/eqP->; rewrite !permE //; case: a; do 4?case.
+Qed.
+
+Lemma rot_is_rot : rot = rotations.
+Proof.
+apply/setP=> r; apply/idP/idP; last by move/rotations_is_rot; rewrite inE.
+rewrite !inE => h.
+have -> : r = r1 ^+ (r c0) by apply: rot_eq_c0; rewrite // -rot_r1.
+have e2: 2 = r2 c0 by rewrite permE /=.
+have e3: 3 = r3 c0 by rewrite permE /=.
+case (r c0); do 4?[case] => // ?; rewrite ?(expg1, eqxx, orbT) //.
+ by rewrite [nat_of_ord _]/= e2 -rot_r1 ?(eqxx, orbT, rotations_is_rot, inE).
+by rewrite [nat_of_ord _]/= e3 -rot_r1 ?(eqxx, orbT, rotations_is_rot, inE).
+Qed.
+
+(*symmetries*)
+Definition Sh (sc : square) : square := tnth [tuple c1; c0; c3; c2] sc.
+
+Lemma Sh_inj : injective Sh.
+Proof.
+by apply:(can_inj (g:= Sh)); case; do 4?case => //=;move=> H;apply /eqP.
+Qed.
+
+Definition sh := (perm Sh_inj).
+
+Lemma sh_inv : sh^-1 = sh.
+Proof.
+apply:(mulIg sh);rewrite mulVg ;apply/permP.
+by case; do 4?case => //=; move=> H;rewrite !permE /= !permE; apply /eqP.
+Qed.
+
+Definition Sv (sc : square) : square := tnth [tuple c3; c2; c1; c0] sc.
+
+Lemma Sv_inj : injective Sv.
+Proof.
+by apply : (can_inj (g:= Sv));case; do 4?case => //=;move => H;apply /eqP.
+Qed.
+
+Definition sv := (perm Sv_inj).
+
+Lemma sv_inv : sv^-1 = sv.
+Proof.
+apply:(mulIg sv);rewrite mulVg ;apply/permP.
+by case; do 4?case => //=; move=> H; rewrite !permE /= !permE; apply /eqP.
+Qed.
+
+Definition Sd1 (sc : square) : square := tnth [tuple c0; c3; c2; c1] sc.
+
+Lemma Sd1_inj : injective Sd1.
+Proof.
+by apply: can_inj Sd1 _; case; do 4?case=> //=; move=> H; apply /eqP.
+Qed.
+
+Definition sd1 := (perm Sd1_inj).
+
+Lemma sd1_inv : sd1^-1 = sd1.
+Proof.
+apply: (mulIg sd1); rewrite mulVg; apply/permP.
+by case; do 4?case=> //=; move=> H; rewrite !permE /= !permE; apply /eqP.
+Qed.
+
+Definition Sd2 (sc : square) : square := tnth [tuple c2; c1; c0; c3] sc.
+
+Lemma Sd2_inj : injective Sd2.
+Proof.
+by apply: can_inj Sd2 _; case; do 4?case=> //=; move=> H; apply /eqP.
+Qed.
+
+Definition sd2 := (perm Sd2_inj).
+
+Lemma sd2_inv : sd2^-1 = sd2.
+Proof.
+apply: (mulIg sd2); rewrite mulVg; apply/permP.
+by case; do 4?case=> //=; move=> H; rewrite !permE /= !permE; apply/eqP.
+Qed.
+
+Lemma ord_enum4 : enum 'I_4 = [:: c0; c1; c2; c3].
+Proof. by apply: (inj_map val_inj); rewrite val_enum_ord. Qed.
+
+Lemma diff_id_sh : 1 != sh.
+Proof.
+by apply/eqP; move/(congr1 (fun p : {perm square} => p c0)); rewrite !permE.
+Qed.
+
+Definition isometries2 := [set 1; sh].
+
+Lemma card_iso2 : #|isometries2| = 2.
+Proof. by rewrite cards2 diff_id_sh. Qed.
+
+Lemma group_set_iso2 : group_set isometries2.
+Proof.
+apply/group_setP; split => [|x y]; rewrite !inE ?eqxx //.
+do 2![case/orP; move/eqP->]; gsimpl; rewrite ?(eqxx, orbT) //.
+by rewrite -/sh -{1}sh_inv mulVg eqxx.
+Qed.
+Canonical iso2_group := Group group_set_iso2.
+
+Definition isometries :=
+ [set p | [|| p == 1, p == r1, p == r2, p == r3,
+ p == sh, p == sv, p == sd1 | p == sd2 ]].
+
+Definition opp (sc : square) := tnth [tuple c2; c3; c0; c1] sc.
+
+Definition is_iso (p : {perm square}) := forall ci, p (opp ci) = opp (p ci).
+
+Lemma isometries_iso : forall p, p \in isometries -> is_iso p.
+Proof.
+move=> p; rewrite inE.
+by do ?case/orP; move/eqP=> -> a; rewrite !permE; case: a; do 4?case.
+Qed.
+
+Ltac non_inj p a1 a2 heq1 heq2 :=
+let h1:= fresh "h1" in
+(absurd (p a1 = p a2);first (by red; move=> h1;move:(perm_inj h1));
+by rewrite heq1 heq2;apply/eqP).
+
+Ltac is_isoPtac p f e0 e1 e2 e3 :=
+ suff ->: p = f by [rewrite inE eqxx ?orbT];
+ let e := fresh "e" in apply/permP;
+ do 5?[case] => // ?; [move: e0 | move: e1 | move: e2 | move: e3] => e;
+ apply: etrans (congr1 p _) (etrans e _); apply/eqP; rewrite // permE.
+
+Lemma is_isoP : forall p, reflect (is_iso p) (p \in isometries).
+Proof.
+move=> p; apply: (iffP idP) => [|iso_p]; first exact: isometries_iso.
+move e1: (p c1) (iso_p c1) => k1; move e0: (p c0) (iso_p c0) k1 e1 => k0.
+case: k0 e0; do 4?[case] => //= ? e0 e2; do 5?[case] => //= ? e1 e3;
+ try by [non_inj p c0 c1 e0 e1 | non_inj p c0 c3 e0 e3].
+by is_isoPtac p id1 e0 e1 e2 e3.
+by is_isoPtac p sd1 e0 e1 e2 e3.
+by is_isoPtac p sh e0 e1 e2 e3.
+by is_isoPtac p r1 e0 e1 e2 e3.
+by is_isoPtac p sd2 e0 e1 e2 e3.
+by is_isoPtac p r2 e0 e1 e2 e3.
+by is_isoPtac p r3 e0 e1 e2 e3.
+by is_isoPtac p sv e0 e1 e2 e3.
+Qed.
+
+
+Lemma group_set_iso : group_set isometries.
+Proof.
+apply/group_setP; split; first by rewrite inE eqxx /=.
+by move=> x y hx hy; apply/is_isoP => ci; rewrite !permM !isometries_iso.
+Qed.
+
+Canonical iso_group := Group group_set_iso.
+
+Lemma card_rot : #|rot| = 4.
+Proof.
+rewrite -[4]/(size [:: id1; r1; r2; r3]) -(card_uniqP _).
+ by apply: eq_card => x; rewrite rot_is_rot !inE -!orbA.
+by apply: map_uniq (fun p : {perm square} => p c0) _ _; rewrite /= !permE.
+Qed.
+
+Lemma group_set_rotations : group_set rotations.
+Proof. by rewrite -rot_is_rot group_set_rot. Qed.
+
+Canonical rotations_group := Group group_set_rotations.
+
+Notation col_squares := {ffun square -> colors}.
+
+Definition act_f (sc : col_squares) (p : {perm square}) : col_squares :=
+ [ffun z => sc (p^-1 z)].
+
+Lemma act_f_1 : forall k, act_f k 1 = k.
+Proof. by move=> k; apply/ffunP=> a; rewrite ffunE invg1 permE. Qed.
+
+Lemma act_f_morph : forall k x y, act_f k (x * y) = act_f (act_f k x) y.
+Proof. by move=> k x y; apply/ffunP=> a; rewrite !ffunE invMg permE. Qed.
+
+Definition to := TotalAction act_f_1 act_f_morph.
+
+Definition square_coloring_number2 := #|orbit to isometries2 @: setT|.
+Definition square_coloring_number4 := #|orbit to rotations @: setT|.
+Definition square_coloring_number8 := #|orbit to isometries @: setT|.
+
+Lemma Fid : 'Fix_to(1) = setT.
+Proof. apply/setP=> x /=; rewrite in_setT; apply/afix1P; exact: act1. Qed.
+
+Lemma card_Fid : #|'Fix_to(1)| = (n ^ 4)%N.
+Proof.
+rewrite -[4]card_ord -[n]card_ord -card_ffun_on Fid cardsE.
+by symmetry; apply: eq_card => f; exact/ffun_onP.
+Qed.
+
+Definition coin0 (sc : col_squares) : colors := sc c0.
+Definition coin1 (sc : col_squares) : colors := sc c1.
+Definition coin2 (sc : col_squares) : colors := sc c2.
+Definition coin3 (sc : col_squares) : colors := sc c3.
+
+Lemma eqperm_map : forall p1 p2 : col_squares,
+ (p1 == p2) = all (fun s => p1 s == p2 s) [:: c0; c1; c2; c3].
+Proof.
+move=> p1 p2; apply/eqP/allP=> [-> // | Ep12]; apply/ffunP=> x.
+by apply/eqP; apply Ep12; case: x; do 4!case=> //.
+Qed.
+
+Lemma F_Sh :
+ 'Fix_to[sh] = [set x | (coin0 x == coin1 x) && (coin2 x == coin3 x)].
+Proof.
+apply/setP=> x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f sh_inv !ffunE !permE /=.
+by rewrite eq_sym (eq_sym (x c3)) andbT andbA !andbb.
+Qed.
+
+Lemma F_Sv :
+ 'Fix_to[sv] = [set x | (coin0 x == coin3 x) && (coin2 x == coin1 x)].
+Proof.
+apply/setP=> x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f sv_inv !ffunE !permE /=.
+by rewrite eq_sym andbT andbC (eq_sym (x c1)) andbA -andbA !andbb andbC.
+Qed.
+
+Ltac inv_tac :=
+ apply: esym (etrans _ (mul1g _)); apply: canRL (mulgK _) _;
+ let a := fresh "a" in apply/permP => a;
+ apply/eqP; rewrite permM !permE; case: a; do 4?case.
+
+Lemma r1_inv : r1^-1 = r3.
+Proof. by inv_tac. Qed.
+
+Lemma r2_inv : r2^-1 = r2.
+Proof. by inv_tac. Qed.
+
+Lemma r3_inv : r3^-1 = r1.
+Proof. by inv_tac. Qed.
+
+Lemma F_r2 :
+ 'Fix_to[r2] = [set x | (coin0 x == coin2 x) && (coin1 x == coin3 x)].
+Proof.
+apply/setP=> x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f r2_inv !ffunE !permE /=.
+by rewrite eq_sym andbT andbCA andbC (eq_sym (x c3)) andbA -andbA !andbb andbC.
+Qed.
+
+Lemma F_r1 : 'Fix_to[r1] =
+ [set x | (coin0 x == coin1 x)&&(coin1 x == coin2 x)&&(coin2 x == coin3 x)].
+Proof.
+apply/setP=> x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f r1_inv !ffunE !permE andbC.
+by do 3![case E: {+}(_ == _); rewrite // {E}(eqP E)]; rewrite eqxx.
+Qed.
+
+Lemma F_r3 : 'Fix_to[r3] =
+ [set x | (coin0 x == coin1 x)&&(coin1 x == coin2 x)&&(coin2 x == coin3 x)].
+Proof.
+apply/setP=> x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f r3_inv !ffunE !permE /=.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite // {E}(eqP E)].
+Qed.
+
+Lemma card_n2 : forall x y z t : square, uniq [:: x; y; z; t] ->
+ #|[set p : col_squares | (p x == p y) && (p z == p t)]| = (n ^ 2)%N.
+Proof.
+move=> x y z t Uxt; rewrite -[n]card_ord.
+pose f (p : col_squares) := (p x, p z); rewrite -(@card_in_image _ _ f).
+ rewrite -mulnn -card_prod; apply: eq_card => [] [c d] /=; apply/imageP.
+ rewrite (cat_uniq [::x; y]) in Uxt; case/and3P: Uxt => _.
+ rewrite /= !orbF !andbT; case/norP; rewrite !inE => nxzt nyzt _.
+ exists [ffun i => if pred2 x y i then c else d].
+ by rewrite inE !ffunE /= !eqxx orbT (negbTE nxzt) (negbTE nyzt) !eqxx.
+ by rewrite {}/f !ffunE /= eqxx (negbTE nxzt).
+move=> p1 p2; rewrite !inE.
+case/andP=> p1y p1t; case/andP=> p2y p2t [px pz].
+have eqp12: all (fun i => p1 i == p2 i) [:: x; y; z; t].
+ by rewrite /= -(eqP p1y) -(eqP p1t) -(eqP p2y) -(eqP p2t) px pz !eqxx.
+apply/ffunP=> i; apply/eqP; apply: (allP eqp12).
+by rewrite (subset_cardP _ (subset_predT _)) // (card_uniqP Uxt) card_ord.
+Qed.
+
+Lemma card_n :
+ #|[set x | (coin0 x == coin1 x)&&(coin1 x == coin2 x)&& (coin2 x == coin3 x)]|
+ = n.
+Proof.
+rewrite -[n]card_ord /coin0 /coin1 /coin2 /coin3.
+pose f (p : col_squares) := p c3; rewrite -(@card_in_image _ _ f).
+ apply: eq_card => c /=; apply/imageP.
+ exists ([ffun => c] : col_squares); last by rewrite /f ffunE.
+ by rewrite /= inE !ffunE !eqxx.
+move=> p1 p2; rewrite /= !inE /f -!andbA => eqp1 eqp2 eqp12.
+apply/eqP; rewrite eqperm_map /= andbT.
+case/and3P: eqp1; do 3!move/eqP->; case/and3P: eqp2; do 3!move/eqP->.
+by rewrite !andbb eqp12.
+Qed.
+
+Lemma burnside_app2 : (square_coloring_number2 * 2 = n ^ 4 + n ^ 2)%N.
+Proof.
+rewrite (burnside_formula [:: id1; sh]) => [||p]; last first.
+- by rewrite !inE.
+- by rewrite /= inE diff_id_sh.
+by rewrite 2!big_cons big_nil addn0 {1}card_Fid F_Sh card_n2.
+Qed.
+
+Lemma burnside_app_rot :
+ (square_coloring_number4 * 4 = n ^ 4 + n ^ 2 + 2 * n)%N.
+Proof.
+rewrite (burnside_formula [:: id1; r1; r2; r3]) => [||p]; last first.
+- by rewrite !inE !orbA.
+- by apply: map_uniq (fun p : {perm square} => p c0) _ _; rewrite /= !permE.
+rewrite !big_cons big_nil /= addn0 {1}card_Fid F_r1 F_r2 F_r3.
+by rewrite card_n card_n2 //=; ring.
+Qed.
+
+Lemma F_Sd1 : 'Fix_to[sd1] = [set x | coin1 x == coin3 x].
+Proof.
+apply/setP => x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+rewrite /act_f sd1_inv !ffunE !permE /=.
+by rewrite !eqxx !andbT eq_sym /= andbb.
+Qed.
+
+Lemma card_n3 : forall x y : square, x != y ->
+ #|[set k : col_squares | k x == k y]| = (n ^ 3)%N.
+Proof.
+move=> x y nxy; apply/eqP; case: (ltngtP n 0) => // [|n0]; last first.
+ by rewrite n0; apply/existsP=> [] [p _]; case: (p c0) => i; rewrite n0.
+move/eqn_pmul2l <-; rewrite -expnS -card_Fid Fid cardsT.
+rewrite -{1}[n]card_ord -cardX.
+pose pk k := [ffun i => k (if i == y then x else i) : colors].
+rewrite -(@card_image _ _ (fun k : col_squares => (k y, pk k))).
+ apply/eqP; apply: eq_card => ck /=; rewrite inE /= [_ \in _]inE.
+ apply/eqP/imageP; last first.
+ by case=> k _ -> /=; rewrite !ffunE if_same eqxx.
+ case: ck => c k /= kxy.
+ exists [ffun i => if i == y then c else k i]; first by rewrite inE.
+ rewrite !ffunE eqxx; congr (_, _); apply/ffunP=> i; rewrite !ffunE.
+ case Eiy: (i == y); last by rewrite Eiy.
+ by rewrite (negbTE nxy) (eqP Eiy).
+move=> k1 k2 [Eky Epk]; apply/ffunP=> i.
+have{Epk}: pk k1 i = pk k2 i by rewrite Epk.
+by rewrite !ffunE; case: eqP => // ->.
+Qed.
+
+Lemma F_Sd2 : 'Fix_to[sd2] = [set x | coin0 x == coin2 x].
+Proof.
+apply/setP => x; rewrite (sameP afix1P eqP) !inE eqperm_map /=.
+by rewrite /act_f sd2_inv !ffunE !permE /= !eqxx !andbT eq_sym /= andbb.
+Qed.
+
+Lemma burnside_app_iso :
+ (square_coloring_number8 * 8 = n ^ 4 + 2 * n ^ 3 + 3 * n ^ 2 + 2 * n)%N.
+Proof.
+pose iso_list := [:: id1; r1; r2; r3; sh; sv; sd1; sd2].
+rewrite (burnside_formula iso_list) => [||p]; last first.
+- by rewrite /= !inE.
+- apply: map_uniq (fun p : {perm square} => (p c0, p c1)) _ _.
+ by rewrite /= !permE.
+rewrite !big_cons big_nil {1}card_Fid F_r1 F_r2 F_r3 F_Sh F_Sv F_Sd1 F_Sd2.
+by rewrite card_n !card_n3 // !card_n2 //=; ring.
+Qed.
+
+End square_colouring.
+
+Section cube_colouring.
+
+Definition cube := 'I_6.
+Canonical cube_eqType := Eval hnf in [eqType of cube].
+Canonical cube_choiceType := Eval hnf in [choiceType of cube].
+Canonical cube_countType := Eval hnf in [countType of cube].
+Canonical cube_finType := Eval hnf in [finType of cube].
+Canonical cube_subType := Eval hnf in [subType of cube].
+Canonical cube_subCountType := Eval hnf in [subCountType of cube].
+Canonical cube_subFinType := Eval hnf in [subFinType of cube].
+
+Definition mkFcube i : cube := Sub (i %% 6) (ltn_mod i 6).
+Definition F0 := mkFcube 0.
+Definition F1 := mkFcube 1.
+Definition F2 := mkFcube 2.
+Definition F3 := mkFcube 3.
+Definition F4 := mkFcube 4.
+Definition F5 := mkFcube 5.
+
+(* axial symetries*)
+Definition S05 := [:: F0; F4; F3; F2; F1; F5].
+Definition S05f (sc : cube) : cube := tnth [tuple of S05] sc.
+
+
+Definition S14 := [:: F5; F1; F3; F2; F4; F0].
+Definition S14f (sc : cube) : cube := tnth [tuple of S14] sc.
+
+Definition S23 := [:: F5; F4; F2; F3; F1; F0].
+Definition S23f (sc : cube) : cube := tnth [tuple of S23] sc.
+
+(* rotations 90 *)
+Definition R05 := [:: F0; F2; F4; F1; F3; F5].
+Definition R05f (sc : cube) : cube := tnth [tuple of R05] sc.
+Definition R50 := [:: F0; F3; F1; F4; F2; F5].
+Definition R50f (sc : cube) : cube := tnth [tuple of R50] sc.
+Definition R14 := [:: F3; F1; F0; F5; F4; F2].
+Definition R14f (sc : cube) : cube := tnth [tuple of R14] sc.
+Definition R41 := [:: F2; F1; F5; F0; F4; F3].
+Definition R41f (sc : cube) : cube := tnth [tuple of R41] sc.
+Definition R23 := [:: F1; F5; F2; F3; F0; F4].
+Definition R23f (sc : cube) : cube := tnth [tuple of R23] sc.
+Definition R32 := [:: F4; F0; F2; F3; F5; F1].
+Definition R32f (sc : cube) : cube := tnth [tuple of R32] sc.
+(* rotations 120 *)
+Definition R024 := [:: F2; F5; F4; F1; F0; F3].
+Definition R024f (sc : cube) : cube := tnth [tuple of R024] sc.
+Definition R042 := [:: F4; F3; F0; F5; F2; F1].
+Definition R042f (sc : cube) : cube := tnth [tuple of R042] sc.
+Definition R012 := [:: F1; F2; F0; F5; F3; F4].
+Definition R012f (sc : cube) : cube := tnth [tuple of R012] sc.
+Definition R021 := [:: F2; F0; F1; F4; F5; F3].
+Definition R021f (sc : cube) : cube := tnth [tuple of R021] sc.
+Definition R031 := [:: F3; F0; F4; F1; F5; F2].
+Definition R031f (sc : cube) : cube := tnth [tuple of R031] sc.
+Definition R013 := [:: F1; F3; F5; F0; F2; F4].
+Definition R013f (sc : cube) : cube := tnth [tuple of R013] sc.
+Definition R043 := [:: F4; F2; F5; F0; F3; F1].
+Definition R043f (sc : cube) : cube := tnth [tuple of R043] sc.
+Definition R034 := [:: F3; F5; F1; F4; F0; F2].
+Definition R034f (sc : cube) : cube := tnth [tuple of R034] sc.
+(* last symmetries*)
+Definition S1 := [:: F5; F2; F1; F4; F3; F0].
+Definition S1f (sc : cube) : cube := tnth [tuple of S1] sc.
+Definition S2 := [:: F5; F3; F4; F1; F2; F0].
+Definition S2f (sc : cube) : cube := tnth [tuple of S2] sc.
+Definition S3 := [:: F1; F0; F3; F2; F5; F4].
+Definition S3f (sc : cube) : cube := tnth [tuple of S3] sc.
+Definition S4 := [:: F4; F5; F3; F2; F0; F1].
+Definition S4f (sc : cube) : cube := tnth [tuple of S4] sc.
+Definition S5 := [:: F2; F4; F0; F5; F1; F3].
+Definition S5f (sc : cube) : cube := tnth [tuple of S5] sc.
+Definition S6 := [::F3; F4; F5; F0; F1; F2].
+Definition S6f (sc : cube) : cube := tnth [tuple of S6] sc.
+
+Lemma S1_inv : involutive S1f.
+Proof. by move=> z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma S2_inv : involutive S2f.
+Proof. by move=> z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma S3_inv : involutive S3f.
+Proof. by move=> z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma S4_inv : involutive S4f.
+Proof. by move=> z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma S5_inv : involutive S5f.
+Proof. by move=> z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma S6_inv : involutive S6f.
+Proof. by move=> z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma S05_inj : injective S05f.
+Proof. by apply: can_inj S05f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma S14_inj : injective S14f.
+Proof. by apply: can_inj S14f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma S23_inv : involutive S23f.
+Proof. by move=> z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R05_inj : injective R05f.
+Proof. by apply: can_inj R50f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R14_inj : injective R14f.
+Proof. by apply: can_inj R41f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R23_inj : injective R23f.
+Proof. by apply: can_inj R32f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R50_inj : injective R50f.
+Proof. by apply: can_inj R05f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R41_inj : injective R41f.
+Proof. by apply: can_inj R14f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R32_inj : injective R32f.
+Proof. by apply: can_inj R23f _ => z; apply /eqP; case: z; do 6?case. Qed.
+
+Lemma R024_inj : injective R024f.
+Proof. by apply: can_inj R042f _ => z; apply /eqP; case: z ; do 6?case. Qed.
+
+Lemma R042_inj : injective R042f.
+Proof. by apply: can_inj R024f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R012_inj : injective R012f.
+Proof. by apply: can_inj R021f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R021_inj : injective R021f.
+Proof. by apply: can_inj R012f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R031_inj : injective R031f.
+Proof. by apply: can_inj R013f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+
+Lemma R013_inj : injective R013f.
+Proof. by apply: can_inj R031f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R043_inj : injective R043f.
+Proof. by apply: can_inj R034f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Lemma R034_inj : injective R034f.
+Proof. by apply: can_inj R043f _ => z; apply/eqP; case: z; do 6?case. Qed.
+
+Definition id3 := 1 : {perm cube}.
+
+Definition s05 := (perm S05_inj).
+
+Definition s14 : {perm cube}.
+Proof.
+apply: (@perm _ S14f); apply: can_inj S14f _ => z.
+by apply /eqP; case: z; do 6?case.
+Defined.
+
+Definition s23 := (perm (inv_inj S23_inv)).
+Definition r05 := (perm R05_inj).
+Definition r14 := (perm R14_inj).
+Definition r23 := (perm R23_inj).
+Definition r50 := (perm R50_inj).
+Definition r41 := (perm R41_inj).
+Definition r32 := (perm R32_inj).
+Definition r024 := (perm R024_inj).
+Definition r042 := (perm R042_inj).
+Definition r012 := (perm R012_inj).
+Definition r021 := (perm R021_inj).
+Definition r031 := (perm R031_inj).
+Definition r013 := (perm R013_inj).
+Definition r043 := (perm R043_inj).
+Definition r034 := (perm R034_inj).
+
+Definition s1 := (perm (inv_inj S1_inv)).
+Definition s2 := (perm (inv_inj S2_inv)).
+Definition s3 := (perm (inv_inj S3_inv)).
+Definition s4 := (perm (inv_inj S4_inv)).
+Definition s5 := (perm (inv_inj S5_inv)).
+Definition s6 := (perm (inv_inj S6_inv)).
+
+Definition dir_iso3 := [set p |
+[|| id3 == p, s05 == p, s14 == p, s23 == p, r05 == p, r14 == p, r23 == p,
+ r50 == p, r41 == p, r32 == p, r024 == p, r042 == p, r012 == p, r021 == p,
+ r031 == p, r013 == p, r043 == p, r034 == p,
+ s1 == p, s2 == p, s3 == p, s4 == p, s5 == p | s6 == p]].
+
+Definition dir_iso3l := [:: id3; s05; s14; s23; r05; r14; r23; r50; r41;
+ r32; r024; r042; r012; r021; r031; r013; r043 ; r034;
+ s1 ; s2; s3; s4; s5; s6].
+
+Definition S0 := [:: F5; F4; F3; F2; F1; F0].
+Definition S0f (sc : cube) : cube := tnth [tuple of S0] sc.
+
+Lemma S0_inv : involutive S0f.
+Proof. by move=> z; apply/eqP; case: z; do 6?case. Qed.
+
+Definition s0 := (perm (inv_inj S0_inv)).
+
+Definition is_iso3 (p : {perm cube}) := forall fi, p (s0 fi) = s0 (p fi).
+
+Lemma dir_iso_iso3 : forall p, p \in dir_iso3 -> is_iso3 p.
+Proof.
+move=> p; rewrite inE.
+by do ?case/orP; move/eqP=> <- a; rewrite !permE; case: a; do 6?case.
+Qed.
+
+Lemma iso3_ndir : forall p, p \in dir_iso3 -> is_iso3 (s0 * p).
+Proof.
+move=> p; rewrite inE.
+by do ?case/orP; move/eqP=> <- a; rewrite !(permM, permE); case: a; do 6?case.
+Qed.
+
+Definition sop (p : {perm cube}) : seq cube := val (val (val p)).
+
+Lemma sop_inj : injective sop.
+Proof. do 2!apply: (inj_comp val_inj); exact: val_inj. Qed.
+
+Definition prod_tuple (t1 t2 : seq cube) :=
+ map (fun n : 'I_6 => nth F0 t2 n) t1.
+
+Lemma sop_spec : forall x (n0 : 'I_6), nth F0 (sop x) n0 = x n0.
+Proof. by move=> x n0; rewrite -pvalE unlock enum_rank_ord (tnth_nth F0). Qed.
+
+Lemma prod_t_correct : forall (x y : {perm cube}) (i : cube),
+ (x * y) i = nth F0 (prod_tuple (sop x) (sop y)) i.
+Proof.
+move=> x y i; rewrite permM -!sop_spec (nth_map F0) // size_tuple /=.
+by rewrite card_ord ltn_ord.
+Qed.
+
+Lemma sop_morph : {morph sop : x y / x * y >-> prod_tuple x y}.
+Proof.
+move=> x y; apply: (@eq_from_nth _ F0) => [|/= i].
+ by rewrite size_map !size_tuple.
+rewrite size_tuple card_ord => lti6.
+by rewrite -[i]/(val (Ordinal lti6)) sop_spec -prod_t_correct.
+Qed.
+
+Definition ecubes : seq cube := [:: F0; F1; F2; F3; F4; F5].
+
+Lemma ecubes_def : ecubes = enum (@predT cube).
+Proof. by apply: (inj_map val_inj); rewrite val_enum_ord. Qed.
+
+Definition seq_iso_L := [::
+ [:: F0; F1; F2; F3; F4; F5];
+ S05; S14; S23; R05; R14; R23; R50; R41; R32;
+ R024; R042; R012; R021; R031; R013; R043; R034;
+ S1; S2; S3; S4; S5; S6].
+
+Lemma seqs1 : forall f injf, sop (@perm _ f injf) = map f ecubes.
+Proof.
+move=> f ?; rewrite ecubes_def /sop /= -codom_ffun pvalE.
+apply: eq_codom; exact: permE.
+Qed.
+
+Lemma Lcorrect : seq_iso_L == map sop [:: id3; s05; s14; s23; r05; r14; r23;
+ r50; r41; r32; r024; r042; r012; r021; r031; r013; r043 ; r034;
+ s1 ; s2; s3; s4; s5; s6].
+Proof. by rewrite /= !seqs1. Qed.
+
+Lemma iso0_1 : dir_iso3 =i dir_iso3l.
+Proof. by move=> p; rewrite /= !inE /= -!(eq_sym p). Qed.
+
+Lemma L_iso : forall p, (p \in dir_iso3) = (sop p \in seq_iso_L).
+Proof.
+move=> p; rewrite (eqP Lcorrect) mem_map ?iso0_1 //; exact: sop_inj.
+Qed.
+
+Lemma stable : forall x y,
+ x \in dir_iso3 -> y \in dir_iso3 -> x * y \in dir_iso3.
+Proof.
+move=> x y; rewrite !L_iso sop_morph => Hx Hy.
+by move/sop: y Hy; apply/allP; move/sop: x Hx; apply/allP; vm_compute.
+Qed.
+
+Lemma iso_eq_F0_F1 : forall r s : {perm cube}, r \in dir_iso3 ->
+ s \in dir_iso3 -> r F0 = s F0 -> r F1 = s F1 -> r = s.
+Proof.
+move=> r s; rewrite !L_iso => hr hs hrs0 hrs1; apply: sop_inj; apply/eqP.
+move/eqP: hrs0; apply/implyP; move/eqP: hrs1; apply/implyP; rewrite -!sop_spec.
+by move/sop: r hr; apply/allP; move/sop: s hs; apply/allP; vm_compute.
+Qed.
+
+Lemma ndir_s0p : forall p, p \in dir_iso3 -> s0 * p \notin dir_iso3.
+Proof.
+move=> p; rewrite !L_iso sop_morph seqs1.
+by move/sop: p; apply/allP; vm_compute.
+Qed.
+
+Definition indir_iso3l := map (mulg s0) dir_iso3l.
+
+Definition iso3l := dir_iso3l ++ indir_iso3l.
+
+Definition seq_iso3_L := map sop iso3l.
+
+Lemma eqperm : forall p1 p2 : {perm cube},
+ (p1 == p2) = all (fun s => p1 s == p2 s) ecubes.
+Proof.
+move=> p1 p2; apply/eqP/allP=> [-> // | Ep12]; apply/permP=> x.
+by apply/eqP; rewrite Ep12 // ecubes_def mem_enum.
+Qed.
+
+Lemma iso_eq_F0_F1_F2 : forall r s : {perm cube}, is_iso3 r ->
+ is_iso3 s -> r F0 = s F0 -> r F1 = s F1 -> r F2 = s F2 -> r = s.
+Proof.
+move=> r s hr hs hrs0 hrs1 hrs2.
+have:= hrs0; have:= hrs1; have:= hrs2.
+have e23: F2 = s0 F3 by apply/eqP; rewrite permE /S0f (tnth_nth F0).
+have e14: F1 = s0 F4 by apply/eqP; rewrite permE /S0f (tnth_nth F0).
+have e05: F0 = s0 F5 by apply/eqP; rewrite permE /S0f (tnth_nth F0).
+rewrite e23 e14 e05; rewrite !hr !hs.
+move/perm_inj=> hrs3; move/perm_inj=> hrs4; move/perm_inj=> hrs5.
+by apply/eqP; rewrite eqperm /= hrs0 hrs1 hrs2 hrs3 hrs4 hrs5 !eqxx.
+Qed.
+
+Ltac iso_tac :=
+ let a := fresh "a" in apply/permP => a;
+ apply/eqP; rewrite !permM !permE; case: a; do 6? case.
+
+Ltac inv_tac :=
+ apply: esym (etrans _ (mul1g _)); apply: canRL (mulgK _) _; iso_tac.
+
+Lemma dir_s0p : forall p, (s0 * p) \in dir_iso3 -> p \notin dir_iso3.
+Proof.
+move => p Hs0p; move: (ndir_s0p Hs0p); rewrite mulgA.
+have e: (s0^-1=s0) by inv_tac.
+by rewrite -{1}e mulVg mul1g.
+Qed.
+
+Definition is_iso3b p := (p * s0 == s0 * p).
+Definition iso3 := [set p | is_iso3b p].
+
+Lemma is_iso3P : forall p, reflect (is_iso3 p) (p \in iso3).
+Proof.
+move => p; apply: (iffP idP); rewrite inE /iso3 /is_iso3b /is_iso3 => e.
+ by move => fi; rewrite -!permM (eqP e).
+by apply/eqP;apply/permP=> z; rewrite !permM (e z).
+Qed.
+
+Lemma group_set_iso3 : group_set iso3.
+Proof.
+apply /group_setP;split.
+ by apply/is_iso3P => fi; rewrite -!permM mulg1 mul1g.
+move => x1 y; rewrite /iso3 !inE /= /is_iso3.
+rewrite /is_iso3b.
+rewrite -mulgA.
+move/eqP => hx1; move/eqP => hy.
+rewrite hy !mulgA. by rewrite -hx1.
+Qed.
+
+Canonical iso_group3 := Group group_set_iso3.
+
+Lemma group_set_diso3 : group_set dir_iso3.
+Proof.
+apply/group_setP;split;first by rewrite inE eqxx /=.
+by exact:stable.
+Qed.
+Canonical diso_group3 := Group group_set_diso3.
+
+Lemma gen_diso3 : dir_iso3 = <<[set r05; r14]>>.
+Proof.
+apply/setP; apply/subset_eqP;apply/andP; split;first last.
+ by rewrite gen_subG;apply/subsetP => x; rewrite !inE;
+ case/orP; move/eqP ->; rewrite eqxx !orbT.
+apply/subsetP => x; rewrite !inE.
+have -> : s05 = r05 * r05 by iso_tac.
+have -> : s14 = r14 * r14 by iso_tac.
+have -> : s23 = r14 * r14 * r05 * r05 by iso_tac.
+have -> : r23 = r05 * r14 * r05 * r14 * r14 by iso_tac.
+have -> : r50 = r05 * r05 * r05 by iso_tac.
+have -> : r41 = r14 * r14 * r14 by iso_tac.
+have -> : r32 = r14 * r14 * r14 * r05* r14 by iso_tac.
+have -> : r024 = r05 * r14 * r14 * r14 by iso_tac.
+have -> : r042 = r14 * r05 * r05 * r05 by iso_tac.
+have -> : r012 = r14 * r05 by iso_tac.
+have -> : r021 = r05 * r14 * r05 * r05 by iso_tac.
+have -> : r031 = r05 * r14 by iso_tac.
+have -> : r013 = r05 * r05 * r14 * r05 by iso_tac.
+have -> : r043 = r14 * r14 * r14 * r05 by iso_tac.
+have -> : r034 = r05 * r05 * r05 * r14 by iso_tac.
+have -> : s1 = r14 * r14 * r05 by iso_tac.
+have -> : s2 = r05 * r14 * r14 by iso_tac.
+have -> : s3 = r05 * r14 * r05 by iso_tac.
+have -> : s4 = r05 * r14 * r14 * r14 * r05 by iso_tac.
+have -> : s5 = r14 * r05 * r05 by iso_tac.
+have -> : s6 = r05 * r05 * r14 by iso_tac.
+do ?case/predU1P=> [<-|]; first exact: group1; last (move/eqP => <-);
+ by rewrite ?groupMl ?mem_gen // !inE eqxx ?orbT.
+Qed.
+
+Notation col_cubes := {ffun cube -> colors}.
+
+Definition act_g (sc : col_cubes) (p : {perm cube}) : col_cubes :=
+ [ffun z => sc (p^-1 z)].
+
+Lemma act_g_1 : forall k, act_g k 1 = k.
+Proof. by move=> k; apply/ffunP=> a; rewrite ffunE invg1 permE. Qed.
+
+Lemma act_g_morph : forall k x y, act_g k (x * y) = act_g (act_g k x) y.
+Proof. by move=> k x y; apply/ffunP=> a; rewrite !ffunE invMg permE. Qed.
+
+Definition to_g := TotalAction act_g_1 act_g_morph.
+
+Definition cube_coloring_number24 := #|orbit to_g diso_group3 @: setT|.
+
+Lemma Fid3 : 'Fix_to_g[1] = setT.
+Proof. by apply/setP=> x /=; rewrite (sameP afix1P eqP) !inE act1 eqxx. Qed.
+
+Lemma card_Fid3 : #|'Fix_to_g[1]| = (n ^ 6)%N.
+Proof.
+rewrite -[6]card_ord -[n]card_ord -card_ffun_on Fid3 cardsT.
+by symmetry; apply: eq_card => ff; exact/ffun_onP.
+Qed.
+
+Definition col0 (sc : col_cubes) : colors := sc F0.
+Definition col1 (sc : col_cubes) : colors := sc F1.
+Definition col2 (sc : col_cubes) : colors := sc F2.
+Definition col3 (sc : col_cubes) : colors := sc F3.
+Definition col4 (sc : col_cubes) : colors := sc F4.
+Definition col5 (sc : col_cubes) : colors := sc F5.
+
+Lemma eqperm_map2 : forall p1 p2 : col_cubes,
+ (p1 == p2) = all (fun s => p1 s == p2 s) [:: F0; F1; F2; F3; F4; F5].
+Proof.
+move=> p1 p2; apply/eqP/allP=> [-> // | Ep12]; apply/ffunP=> x.
+by apply/eqP; apply Ep12; case: x; do 6?case.
+Qed.
+
+Notation infE := (sameP afix1P eqP).
+
+Lemma F_s05 :
+ 'Fix_to_g[s05] = [set x | (col1 x == col4 x) && (col2 x == col3 x)].
+Proof.
+have s05_inv: s05^-1=s05 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s05_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= andbT/col1/col2/col3/col4/col5/col0.
+by do 2![rewrite eq_sym; case : {+}(_ == _)=> //= ].
+Qed.
+
+Lemma F_s14 :
+ 'Fix_to_g[s14]= [set x | (col0 x == col5 x) && (col2 x == col3 x)].
+Proof.
+have s14_inv: s14^-1=s14 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s14_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= andbT/col1/col2/col3/col4/col5/col0.
+by do 2![rewrite eq_sym; case : {+}(_ == _)=> //= ].
+Qed.
+
+Lemma r05_inv : r05^-1 = r50.
+Proof. by inv_tac. Qed.
+
+Lemma r50_inv : r50^-1 = r05.
+Proof. by inv_tac. Qed.
+
+Lemma r14_inv : r14^-1 = r41.
+Proof. by inv_tac. Qed.
+
+Lemma r41_inv : r41^-1 = r14.
+Proof. by inv_tac. Qed.
+
+Lemma s23_inv : s23^-1 = s23.
+Proof. by inv_tac. Qed.
+
+Lemma F_s23 :
+ 'Fix_to_g[s23] = [set x | (col0 x == col5 x) && (col1 x == col4 x)].
+Proof.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s23_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= andbT/col1/col2/col3/col4/col5/col0.
+by do 2![rewrite eq_sym; case : {+}(_ == _)=> //=].
+Qed.
+
+Lemma F_r05 : 'Fix_to_g[r05]=
+ [set x | (col1 x == col2 x) && (col2 x == col3 x)
+ && (col3 x == col4 x)].
+Proof.
+apply sym_equal.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r05_inv !ffunE !permE /=.
+rewrite !eqxx /= !andbT /col1/col2/col3/col4/col5/col0.
+by do 3! [rewrite eq_sym;case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E) ].
+Qed.
+
+Lemma F_r50 : 'Fix_to_g[r50]=
+ [set x | (col1 x == col2 x) && (col2 x == col3 x)
+ && (col3 x == col4 x)].
+Proof.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r50_inv !ffunE !permE /=.
+apply sym_equal;rewrite !eqxx /= !andbT /col1/col2/col3/col4.
+by do 3![rewrite eq_sym;case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E) ].
+Qed.
+
+Lemma F_r23 : 'Fix_to_g[r23] =
+ [set x | (col0 x == col1 x) && (col1 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r23_inv: r23^-1 = r32 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r23_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= !andbT /col1/col0/col5/col4.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E)].
+Qed.
+
+Lemma F_r32 : 'Fix_to_g[r32] =
+ [set x | (col0 x == col1 x) && (col1 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r32_inv: r32^-1 = r23 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r32_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= !andbT /col1/col0/col5/col4.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E)].
+Qed.
+
+Lemma F_r14 : 'Fix_to_g[r14] =
+ [set x | (col0 x == col2 x) && (col2 x == col3 x) && (col3 x == col5 x)].
+Proof.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r14_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= !andbT /col2/col0/col5/col3.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E)].
+Qed.
+
+Lemma F_r41 : 'Fix_to_g[r41] =
+ [set x | (col0 x == col2 x) && (col2 x == col3 x) && (col3 x == col5 x)].
+Proof.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r41_inv !ffunE !permE /=.
+apply sym_equal; rewrite !eqxx /= !andbT /col2/col0/col5/col3.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // {E}(eqP E)].
+Qed.
+
+Lemma F_r024 : 'Fix_to_g[r024] =
+ [set x | (col0 x == col4 x) && (col4 x == col2 x) && (col1 x == col3 x)
+ && (col3 x == col5 x) ].
+Proof.
+have r024_inv: r024^-1 = r042 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r024_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r042 : 'Fix_to_g[r042] =
+ [set x | (col0 x == col4 x) && (col4 x == col2 x) && (col1 x == col3 x)
+ && (col3 x == col5 x)].
+Proof.
+have r042_inv: r042^-1 = r024 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r042_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r012 : 'Fix_to_g[r012] =
+ [set x | (col0 x == col2 x) && (col2 x == col1 x) && (col3 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r012_inv: r012^-1 = r021 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r012_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r021 : 'Fix_to_g[r021] =
+ [set x | (col0 x == col2 x) && (col2 x == col1 x) && (col3 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r021_inv: r021^-1 = r012 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r021_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r031 : 'Fix_to_g[r031] =
+ [set x | (col0 x == col3 x) && (col3 x == col1 x) && (col2 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r031_inv: r031^-1 = r013 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r031_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r013 : 'Fix_to_g[r013] =
+ [set x | (col0 x == col3 x) && (col3 x == col1 x) && (col2 x == col4 x)
+ && (col4 x == col5 x)].
+Proof.
+have r013_inv: r013^-1 = r031 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r013_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r043 : 'Fix_to_g[r043] =
+ [set x | (col0 x == col4 x) && (col4 x == col3 x) && (col1 x == col2 x)
+ && (col2 x == col5 x)].
+Proof.
+have r043_inv: r043^-1 = r034 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r043_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_r034 : 'Fix_to_g[r034] =
+ [set x | (col0 x == col4 x) && (col4 x == col3 x) && (col1 x == col2 x)
+ && (col2 x == col5 x)].
+Proof.
+have r034_inv: r034^-1 = r043 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g r034_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 4![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s1 : 'Fix_to_g[s1] =
+ [set x | (col0 x == col5 x) && (col1 x == col2 x) && (col3 x == col4 x)].
+Proof.
+have s1_inv: s1^-1 = s1 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s1_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s2 : 'Fix_to_g[s2] =
+ [set x | (col0 x == col5 x) && (col1 x == col3 x) && (col2 x == col4 x)].
+Proof.
+have s2_inv: s2^-1 = s2 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s2_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s3 : 'Fix_to_g[s3] =
+ [set x | (col0 x == col1 x) && (col2 x == col3 x) && (col4 x == col5 x)].
+Proof.
+have s3_inv: s3^-1 = s3 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s3_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s4 : 'Fix_to_g[s4] =
+ [set x | (col0 x == col4 x) && (col1 x == col5 x) && (col2 x == col3 x)].
+Proof.
+have s4_inv: s4^-1 = s4 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s4_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s5 : 'Fix_to_g[s5] =
+ [set x | (col0 x == col2 x) && (col1 x == col4 x) && (col3 x == col5 x)].
+Proof.
+have s5_inv: s5^-1 = s5 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s5_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma F_s6 : 'Fix_to_g[s6] =
+ [set x | (col0 x == col3 x) && (col1 x == col4 x) && (col2 x == col5 x)].
+Proof.
+have s6_inv: s6^-1 = s6 by inv_tac.
+apply/setP => x; rewrite infE !inE eqperm_map2 /= /act_g s6_inv !ffunE !permE /=.
+apply sym_equal; rewrite ?eqxx /= !andbT /col0/col1/col2/col3/col4/col5.
+by do 3![rewrite eq_sym; case E: {+}(_ == _); rewrite ?andbF // ?{E}(eqP E)].
+Qed.
+
+Lemma uniq4_uniq6 : forall x y z t : cube,
+ uniq [:: x; y; z; t] -> exists u, exists v, uniq [:: x; y; z; t; u; v].
+Proof.
+move => x y z t Uxt; move:( cardC (mem [:: x; y; z; t])).
+rewrite card_ord (card_uniq_tuple Uxt) => hcard.
+have hcard2: #|predC (mem [:: x; y; z; t])| = 2.
+ by apply:( @addnI 4); rewrite /injective hcard.
+have: #|predC (mem [:: x; y; z; t])| != 0 by rewrite hcard2.
+case/existsP=> u Hu; exists u.
+move: (cardC (mem [:: x; y; z; t; u])); rewrite card_ord => hcard5.
+have: #|[predC [:: x; y; z; t; u]]| !=0.
+ rewrite -lt0n -(ltn_add2l #|[:: x; y; z; t; u]|) hcard5 addn0.
+ by apply: (leq_ltn_trans (card_size [:: x; y; z; t; u])).
+case/existsP=> v; rewrite inE (mem_cat _ [:: _; _; _; _]).
+case/norP=> Hv Huv; exists v.
+rewrite (cat_uniq [:: x; y; z; t]) Uxt andTb.
+by rewrite -rev_uniq /= negb_or Hu orbF Hv Huv.
+Qed.
+
+Lemma card_n4 : forall x y z t : cube, uniq [:: x; y; z; t] ->
+ #|[set p : col_cubes | (p x == p y) && (p z == p t)]| = (n ^ 4)%N.
+Proof.
+move=> x y z t Uxt. rewrite -[n]card_ord .
+case:(uniq4_uniq6 Uxt) => u; case => v Uxv.
+pose ff (p : col_cubes) := (p x, p z, p u , p v).
+rewrite -(@card_in_image _ _ ff); first last.
+ move=> p1 p2; rewrite !inE.
+ case/andP=> p1y p1t; case/andP=> p2y p2t [px pz] pu pv.
+ have eqp12: all (fun i => p1 i == p2 i) [:: x; y; z; t; u ; v].
+ by rewrite /= -(eqP p1y) -(eqP p1t) -(eqP p2y) -(eqP p2t) px pz pu pv !eqxx.
+ apply/ffunP=> i; apply/eqP; apply: (allP eqp12).
+ by rewrite (subset_cardP _ (subset_predT _)) // (card_uniqP Uxv) card_ord.
+have ->:forall n, (n ^ 4)%N= (n*n*n*n)%N.
+ by move => n0;rewrite (expnD n0 2 2) -mulnn mulnA.
+rewrite -!card_prod; apply: eq_card => [] [[[c d]e ]g] /=; apply/imageP.
+rewrite (cat_uniq [::x; y;z;t]) in Uxv; case/and3P: Uxv => _ hasxt.
+rewrite /= !inE andbT.
+move/negbTE=> nuv .
+rewrite (cat_uniq [::x; y]) in Uxt; case/and3P: Uxt => _.
+rewrite /= !andbT orbF; case/norP; rewrite !inE => nxyz nxyt _.
+move:hasxt; rewrite /= !orbF; case/norP; rewrite !inE orbA.
+case/norP => nxyu nztu.
+rewrite orbA;case/norP=> nxyv nztv.
+exists [ffun i => if pred2 x y i then c else if pred2 z t i then d
+ else if u==i then e else g].
+ rewrite !inE /= !ffunE //= !eqxx orbT //= !eqxx /= orbT.
+ by rewrite (negbTE nxyz) (negbTE nxyt).
+rewrite {}/ff !ffunE /= !eqxx /=.
+rewrite (negbTE nxyz) (negbTE nxyu) (negbTE nztu) (negbTE nxyv) (negbTE nztv).
+by rewrite nuv.
+Qed.
+
+Lemma card_n3_3 : forall x y z t: cube, uniq [:: x; y; z;t] ->
+ #|[set p : col_cubes | (p x == p y) && (p y == p z)&& (p z == p t)]|
+ = (n ^ 3)%N.
+Proof.
+move=> x y z t Uxt; rewrite -[n]card_ord .
+case:(uniq4_uniq6 Uxt) => u; case => v Uxv.
+pose ff (p : col_cubes) := (p x, p u , p v);
+ rewrite -(@card_in_image _ _ ff); first last.
+ move=> p1 p2; rewrite !inE.
+ case/andP ;case/andP => p1xy p1yz p1zt.
+ case/andP ;case/andP => p2xy p2yz p2zt [px pu] pv.
+ have eqp12: all (fun i => p1 i == p2 i) [:: x; y; z; t; u ; v].
+ by rewrite /= -(eqP p1zt) -(eqP p2zt) -(eqP p1yz) -(eqP p2yz) -(eqP p1xy)
+ -(eqP p2xy) px pu pv !eqxx.
+ apply/ffunP=> i; apply/eqP; apply: (allP eqp12).
+ by rewrite (subset_cardP _ (subset_predT _)) // (card_uniqP Uxv) card_ord.
+have ->:forall n, (n ^ 3)%N= (n*n*n)%N.
+ by move => n0 ; rewrite (expnD n0 2 1) -mulnn expn1.
+rewrite -!card_prod; apply: eq_card => [] [[c d]e ] /=; apply/imageP.
+rewrite (cat_uniq [::x; y;z;t]) in Uxv; case/and3P: Uxv => _ hasxt.
+rewrite /uniq !inE !andbT; move/negbTE=> nuv.
+exists
+ [ffun i => if (i \in [:: x; y; z; t]) then c else if u == i then d else e].
+ by rewrite /= !inE !ffunE !inE !eqxx !orbT !eqxx.
+rewrite {}/ff !ffunE !inE /= !eqxx /=; move: hasxt; rewrite nuv.
+by do 8![case E: ( _ == _ ); rewrite ?(eqP E)/= ?inE ?eqxx //= ?E {E}].
+Qed.
+
+Lemma card_n2_3 : forall x y z t u v: cube, uniq [:: x; y; z;t; u ; v] ->
+ #|[set p : col_cubes | (p x == p y) && (p y == p z)&& (p t == p u )
+ && (p u== p v)]| = (n ^ 2)%N.
+Proof.
+move=> x y z t u v Uxv; rewrite -[n]card_ord .
+pose ff (p : col_cubes) := (p x, p t); rewrite -(@card_in_image _ _ ff); first last.
+ move=> p1 p2; rewrite !inE.
+ case/andP ;case/andP ; case/andP => p1xy p1yz p1tu p1uv.
+ case/andP ;case/andP; case/andP => p2xy p2yz p2tu p2uv [px pu].
+ have eqp12: all (fun i => p1 i == p2 i) [:: x; y; z; t; u ; v].
+ by rewrite /= -(eqP p1yz) -(eqP p2yz) -(eqP p1xy) -(eqP p2xy) -(eqP p1uv)
+ -(eqP p2uv) -(eqP p1tu) -(eqP p2tu) px pu !eqxx.
+ apply/ffunP=> i; apply/eqP; apply: (allP eqp12).
+ by rewrite (subset_cardP _ (subset_predT _)) // (card_uniqP Uxv) card_ord.
+have ->:forall n, (n ^ 2)%N= (n*n)%N by move => n0 ; rewrite -mulnn .
+ rewrite -!card_prod; apply: eq_card => [] [c d]/=; apply/imageP.
+rewrite (cat_uniq [::x; y;z]) in Uxv; case/and3P: Uxv => Uxt hasxt nuv .
+move:hasxt;rewrite /= !orbF; case/norP; rewrite !inE => nxyzt.
+case/norP => nxyzu nxyzv.
+exists [ffun i => if (i \in [:: x; y; z] ) then c else d].
+ rewrite !inE /= !ffunE !inE //= !eqxx !orbT !eqxx //=.
+ by rewrite (negbTE nxyzt) (negbTE nxyzu)(negbTE nxyzv) !eqxx.
+by rewrite {}/ff !ffunE !inE /= !eqxx /= (negbTE nxyzt).
+Qed.
+
+Lemma card_n3s : forall x y z t u v: cube, uniq [:: x; y; z;t; u ; v] ->
+ #|[set p : col_cubes | (p x == p y) && (p z == p t)&& (p u == p v )]|
+ = (n ^ 3)%N.
+Proof.
+move=> x y z t u v Uxv; rewrite -[n]card_ord .
+pose ff (p : col_cubes) := (p x, p z, p u).
+rewrite -(@card_in_image _ _ ff); first last.
+ move=> p1 p2; rewrite !inE.
+ case/andP ;case/andP => p1xy p1zt p1uv.
+ case/andP ;case/andP => p2xy p2zt p2uv [px pz] pu.
+ have eqp12: all (fun i => p1 i == p2 i) [:: x; y; z; t; u ; v].
+ by rewrite /= -(eqP p1xy) -(eqP p2xy) -(eqP p1zt) -(eqP p2zt) -(eqP p1uv)
+ -(eqP p2uv) px pz pu !eqxx.
+ apply/ffunP=> i; apply/eqP; apply: (allP eqp12).
+ by rewrite (subset_cardP _ (subset_predT _)) // (card_uniqP Uxv) card_ord.
+have ->:forall n, (n ^ 3)%N= (n*n*n)%N.
+ by move => n0 ; rewrite (expnD n0 2 1) -mulnn expn1.
+rewrite -!card_prod. apply: eq_card => [] [[c d]e ] /=; apply/imageP.
+rewrite (cat_uniq [::x; y;z;t]) in Uxv; case/and3P: Uxv => Uxt hasxt nuv .
+rewrite (cat_uniq [::x; y]) in Uxt; case/and3P: Uxt => _.
+rewrite /= !orbF !andbT; case/norP ; rewrite !inE => nxyz nxyt _.
+move: hasxt; rewrite /= !orbF; case/norP; rewrite !inE orbA.
+case/norP => nxyu nztu.
+rewrite orbA;case/norP=> nxyv nztv.
+exists [ffun i => if (i \in [:: x; y] ) then c else if (i \in [:: z; t] )
+ then d else e].
+ rewrite !inE /= !ffunE !inE // !eqxx !orbT !eqxx //=.
+ by rewrite (negbTE nxyz) (negbTE nxyt)(negbTE nxyu) (negbTE nztu)
+ (negbTE nxyv) (negbTE nztv) !eqxx.
+rewrite {}/ff !ffunE !inE /= !eqxx /=.
+by rewrite (negbTE nxyz) (negbTE nxyu) (negbTE nztu).
+Qed.
+
+Lemma burnside_app_iso3 :
+ (cube_coloring_number24 * 24 =
+ n ^ 6 + 6 * n ^ 3 + 3 * n ^ 4 + 8 * (n ^ 2) + 6 * n ^ 3)%N.
+Proof.
+pose iso_list :=[::id3; s05; s14; s23; r05; r14; r23; r50; r41; r32;
+ r024; r042; r012; r021; r031; r013; r043 ; r034;
+ s1 ; s2; s3; s4; s5; s6].
+rewrite (burnside_formula iso_list) => [||p]; last first.
+- by rewrite !inE /= !(eq_sym _ p).
+- apply: map_uniq (fun p : {perm cube} => (p F0, p F1)) _ _.
+ have bsr:(fun p : {perm cube} => (p F0, p F1)) =1
+ (fun p => (nth F0 p F0, nth F0 p F1)) \o sop.
+ by move => x; rewrite /= -2!sop_spec.
+ by rewrite (eq_map bsr) map_comp -(eqP Lcorrect); vm_compute.
+rewrite !big_cons big_nil {1}card_Fid3 /= F_s05 F_s14 F_s23 F_r05 F_r14 F_r23
+ F_r50 F_r41 F_r32 F_r024 F_r042 F_r012 F_r021 F_r031 F_r013 F_r043 F_r034
+ F_s1 F_s2 F_s3 F_s4 F_s5 F_s6.
+by rewrite !card_n4 // !card_n3_3 // !card_n2_3 // !card_n3s //; ring.
+Qed.
+
+End cube_colouring.
+
+End colouring.
+
+Corollary burnside_app_iso_3_3col: cube_coloring_number24 3 = 57.
+Proof.
+by apply/eqP; rewrite -(@eqn_pmul2r 24) // burnside_app_iso3.
+Qed.
+
+
+Corollary burnside_app_iso_2_4col: square_coloring_number8 4 = 55.
+Proof. by apply/eqP; rewrite -(@eqn_pmul2r 8) // burnside_app_iso. Qed.
+
+
+
diff --git a/mathcomp/solvable/center.v b/mathcomp/solvable/center.v
new file mode 100644
index 0000000..0ed9200
--- /dev/null
+++ b/mathcomp/solvable/center.v
@@ -0,0 +1,652 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype bigop.
+Require Import finset fingroup morphism perm automorphism quotient action.
+Require Import gproduct gfunctor cyclic.
+
+(******************************************************************************)
+(* Definition of the center of a group and of external central products: *)
+(* 'Z(G) == the center of the group G, i.e., 'C_G(G). *)
+(* cprod_by isoZ == the finGroupType for the central product of H and K *)
+(* with centers identified by the isomorphism gz on 'Z(H); *)
+(* here isoZ : isom 'Z(H) 'Z(K) gz. Note that the actual *)
+(* central product is [set: cprod_by isoZ]. *)
+(* cpairg1 isoZ == the isomorphism from H to cprod_by isoZ, isoZ as above. *)
+(* cpair1g isoZ == the isomorphism from K to cprod_by isoZ, isoZ as above. *)
+(* xcprod H K == the finGroupType for the external central product of H *)
+(* and K with identified centers, provided the dynamically *)
+(* tested condition 'Z(H) \isog 'Z(K) holds. *)
+(* ncprod H n == the finGroupType for the central product of n copies of *)
+(* H with their centers identified; [set: ncprod H 0] is *)
+(* isomorphic to 'Z(H). *)
+(* xcprodm cf eqf == the morphism induced on cprod_by isoZ, where as above *)
+(* isoZ : isom 'Z(H) 'Z(K) gz, by fH : {morphism H >-> rT} *)
+(* and fK : {morphism K >-> rT}, given both *)
+(* cf : fH @* H \subset 'C(fK @* K) and *)
+(* eqf : {in 'Z(H), fH =1 fK \o gz}. *)
+(* Following Aschbacher, we only provide external central products with *)
+(* identified centers, as these are well defined provided the local center *)
+(* isomorphism group of one of the subgroups is full. Nevertheless the *)
+(* entire construction could be carried out under the weaker assumption that *)
+(* gz is an isomorphism between subgroups of 'Z(H) and 'Z(K), and even the *)
+(* uniqueness theorem holds under the weaker assumption that gz map 'Z(H) to *)
+(* a characteristic subgroup of 'Z(K) not isomorphic to any other subgroup of *)
+(* 'Z(K), a condition that holds for example when K is cyclic, as in the *)
+(* structure theorem for p-groups of symplectic type. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Defs.
+
+Variable gT : finGroupType.
+
+Definition center (A : {set gT}) := 'C_A(A).
+
+Canonical center_group (G : {group gT}) : {group gT} :=
+ Eval hnf in [group of center G].
+
+End Defs.
+
+Arguments Scope center [_ group_scope].
+Notation "''Z' ( A )" := (center A) : group_scope.
+Notation "''Z' ( H )" := (center_group H) : Group_scope.
+
+Lemma morphim_center : GFunctor.pcontinuous center.
+Proof. move=> gT rT G D f; exact: morphim_subcent. Qed.
+
+Canonical center_igFun := [igFun by fun _ _ => subsetIl _ _ & morphim_center].
+Canonical center_gFun := [gFun by morphim_center].
+Canonical center_pgFun := [pgFun by morphim_center].
+
+Section Center.
+
+Variables gT : finGroupType.
+Implicit Type rT : finGroupType.
+Implicit Types (x y : gT) (A B : {set gT}) (G H K D : {group gT}).
+
+Lemma subcentP A B x : reflect (x \in A /\ centralises x B) (x \in 'C_A(B)).
+Proof.
+rewrite inE. case: (x \in A); last by right; case.
+by apply: (iffP centP) => [|[]].
+Qed.
+
+Lemma subcent_sub A B : 'C_A(B) \subset 'N_A(B).
+Proof. by rewrite setIS ?cent_sub. Qed.
+
+Lemma subcent_norm G B : 'N_G(B) \subset 'N('C_G(B)).
+Proof. by rewrite normsI ?subIset ?normG // orbC cent_norm. Qed.
+
+Lemma subcent_normal G B : 'C_G(B) <| 'N_G(B).
+Proof. by rewrite /normal subcent_sub subcent_norm. Qed.
+
+Lemma subcent_char G H K : H \char G -> K \char G -> 'C_H(K) \char G.
+Proof.
+case/charP=> sHG chHG /charP[sKG chKG]; apply/charP.
+split=> [|f injf Gf]; first by rewrite subIset ?sHG.
+by rewrite injm_subcent ?chHG ?chKG.
+Qed.
+
+Lemma centerP A x : reflect (x \in A /\ centralises x A) (x \in 'Z(A)).
+Proof. exact: subcentP. Qed.
+
+Lemma center_sub A : 'Z(A) \subset A.
+Proof. exact: subsetIl. Qed.
+
+Lemma center1 : 'Z(1) = [1 gT].
+Proof. by apply/eqP; rewrite eqEsubset center_sub sub1G. Qed.
+
+Lemma centerC A : {in A, centralised 'Z(A)}.
+Proof. by apply/centsP; rewrite centsC subsetIr. Qed.
+
+Lemma center_normal G : 'Z(G) <| G.
+Proof. exact: gFnormal. Qed.
+
+Lemma sub_center_normal H G : H \subset 'Z(G) -> H <| G.
+Proof. by rewrite subsetI centsC /normal => /andP[-> /cents_norm]. Qed.
+
+Lemma center_abelian G : abelian 'Z(G).
+Proof. by rewrite /abelian subIset // centsC subIset // subxx orbT. Qed.
+
+Lemma center_char G : 'Z(G) \char G.
+Proof. exact: gFchar. Qed.
+
+Lemma center_idP A : reflect ('Z(A) = A) (abelian A).
+Proof. exact: setIidPl. Qed.
+
+Lemma center_class_formula G :
+ #|G| = #|'Z(G)| + \sum_(xG in [set x ^: G | x in G :\: 'C(G)]) #|xG|.
+Proof.
+by rewrite acts_sum_card_orbit ?cardsID // astabsJ normsD ?norms_cent ?normG.
+Qed.
+
+Lemma subcent1P A x y : reflect (y \in A /\ commute x y) (y \in 'C_A[x]).
+Proof.
+rewrite inE; case: (y \in A); last by right; case.
+by apply: (iffP cent1P) => [|[]].
+Qed.
+
+Lemma subcent1_id x G : x \in G -> x \in 'C_G[x].
+Proof. move=> Gx; rewrite inE Gx; exact/cent1P. Qed.
+
+Lemma subcent1_sub x G : 'C_G[x] \subset G.
+Proof. exact: subsetIl. Qed.
+
+Lemma subcent1C x y G : x \in G -> y \in 'C_G[x] -> x \in 'C_G[y].
+Proof. by move=> Gx /subcent1P[_ cxy]; exact/subcent1P. Qed.
+
+Lemma subcent1_cycle_sub x G : x \in G -> <[x]> \subset 'C_G[x].
+Proof. by move=> Gx; rewrite cycle_subG ?subcent1_id. Qed.
+
+Lemma subcent1_cycle_norm x G : 'C_G[x] \subset 'N(<[x]>).
+Proof. by rewrite cents_norm // cent_gen cent_set1 subsetIr. Qed.
+
+Lemma subcent1_cycle_normal x G : x \in G -> <[x]> <| 'C_G[x].
+Proof.
+by move=> Gx; rewrite /normal subcent1_cycle_norm subcent1_cycle_sub.
+Qed.
+
+(* Gorenstein. 1.3.4 *)
+Lemma cyclic_center_factor_abelian G : cyclic (G / 'Z(G)) -> abelian G.
+Proof.
+case/cyclicP=> a Ga; case: (cosetP a) => /= z Nz def_a.
+have G_Zz: G :=: 'Z(G) * <[z]>.
+ rewrite -quotientK ?cycle_subG ?quotient_cycle //=.
+ by rewrite -def_a -Ga quotientGK // center_normal.
+rewrite G_Zz abelianM cycle_abelian center_abelian centsC /= G_Zz.
+by rewrite subIset ?centS ?orbT ?mulG_subr.
+Qed.
+
+Lemma cyclic_factor_abelian H G :
+ H \subset 'Z(G) -> cyclic (G / H) -> abelian G.
+Proof.
+move=> sHZ cycGH; apply: cyclic_center_factor_abelian.
+have nG: G \subset 'N(_) := normal_norm (sub_center_normal _).
+have [f <-]:= homgP (homg_quotientS (nG _ sHZ) (nG _ (subxx _)) sHZ).
+exact: morphim_cyclic.
+Qed.
+
+Section Injm.
+
+Variables (rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+
+Hypothesis injf : 'injm f.
+
+Lemma injm_center G : G \subset D -> f @* 'Z(G) = 'Z(f @* G).
+Proof. exact: injm_subcent. Qed.
+
+End Injm.
+
+End Center.
+
+Implicit Arguments center_idP [gT A].
+
+Lemma isog_center (aT rT : finGroupType) (G : {group aT}) (H : {group rT}) :
+ G \isog H -> 'Z(G) \isog 'Z(H).
+Proof. exact: gFisog. Qed.
+
+Section Product.
+
+Variable gT : finGroupType.
+Implicit Types (A B C : {set gT}) (G H K : {group gT}).
+
+Lemma center_prod H K : K \subset 'C(H) -> 'Z(H) * 'Z(K) = 'Z(H * K).
+Proof.
+move=> cHK; apply/setP=> z; rewrite {3}/center centM !inE.
+have cKH: H \subset 'C(K) by rewrite centsC.
+apply/imset2P/and3P=> [[x y /setIP[Hx cHx] /setIP[Ky cKy] ->{z}]| []].
+ by rewrite mem_imset2 ?groupM // ?(subsetP cHK) ?(subsetP cKH).
+case/imset2P=> x y Hx Ky ->{z}.
+rewrite groupMr => [cHx|]; last exact: subsetP Ky.
+rewrite groupMl => [cKy|]; last exact: subsetP Hx.
+by exists x y; rewrite ?inE ?Hx ?Ky.
+Qed.
+
+Lemma center_cprod A B G : A \* B = G -> 'Z(A) \* 'Z(B) = 'Z(G).
+Proof.
+case/cprodP => [[H K -> ->] <- cHK].
+rewrite cprodE ?center_prod //= subIset ?(subset_trans cHK) //.
+by rewrite centS ?center_sub.
+Qed.
+
+Lemma center_bigcprod I r P (F : I -> {set gT}) G :
+ \big[cprod/1]_(i <- r | P i) F i = G ->
+ \big[cprod/1]_(i <- r | P i) 'Z(F i) = 'Z(G).
+Proof.
+elim/big_ind2: _ G => [_ <-|A B C D IHA IHB G dG|_ _ G ->]; rewrite ?center1 //.
+case/cprodP: dG IHA IHB (dG) => [[H K -> ->] _ _] IHH IHK dG.
+by rewrite (IHH H) // (IHK K) // (center_cprod dG).
+Qed.
+
+Lemma cprod_center_id G : G \* 'Z(G) = G.
+Proof. by rewrite cprodE ?subsetIr // mulGSid ?center_sub. Qed.
+
+Lemma center_dprod A B G : A \x B = G -> 'Z(A) \x 'Z(B) = 'Z(G).
+Proof.
+case/dprodP=> [[H1 H2 -> ->] defG cH12 trH12].
+move: defG; rewrite -cprodE // => /center_cprod/cprodP[_ /= <- cZ12].
+by apply: dprodE; rewrite //= setIAC setIA -setIA trH12 (setIidPl _) ?sub1G.
+Qed.
+
+Lemma center_bigdprod I r P (F: I -> {set gT}) G :
+ \big[dprod/1]_(i <- r | P i) F i = G ->
+ \big[dprod/1]_(i <- r | P i) 'Z(F i) = 'Z(G).
+Proof.
+elim/big_ind2: _ G => [_ <-|A B C D IHA IHB G dG|_ _ G ->]; rewrite ?center1 //.
+case/dprodP: dG IHA IHB (dG) => [[H K -> ->] _ _ _] IHH IHK dG.
+by rewrite (IHH H) // (IHK K) // (center_dprod dG).
+Qed.
+
+Lemma Aut_cprod_full G H K :
+ H \* K = G -> 'Z(H) = 'Z(K) ->
+ Aut_in (Aut H) 'Z(H) \isog Aut 'Z(H) ->
+ Aut_in (Aut K) 'Z(K) \isog Aut 'Z(K) ->
+ Aut_in (Aut G) 'Z(G) \isog Aut 'Z(G).
+Proof.
+move=> defG eqZHK; have [_ defHK cHK] := cprodP defG.
+have defZ: 'Z(G) = 'Z(H) by rewrite -defHK -center_prod // eqZHK mulGid.
+have ziHK: H :&: K = 'Z(K).
+ by apply/eqP; rewrite eqEsubset subsetI -{1 2}eqZHK !center_sub setIS.
+have AutZP := Aut_sub_fullP (@center_sub gT _).
+move/AutZP=> AutZHfull /AutZP AutZKfull; apply/AutZP=> g injg gZ.
+have [gH [def_gH ker_gH _ im_gH]] := domP g defZ.
+have [gK [def_gK ker_gK _ im_gK]] := domP g (etrans defZ eqZHK).
+have [injgH injgK]: 'injm gH /\ 'injm gK by rewrite ker_gH ker_gK.
+have [gHH gKK]: gH @* 'Z(H) = 'Z(H) /\ gK @* 'Z(K) = 'Z(K).
+ by rewrite im_gH im_gK -eqZHK -defZ.
+have [|fH [injfH im_fH fHZ]] := AutZHfull gH injgH.
+ by rewrite im_gH /= -defZ.
+have [|fK [injfK im_fK fKZ]] := AutZKfull gK injgK.
+ by rewrite im_gK /= -eqZHK -defZ.
+have cfHK: fK @* K \subset 'C(fH @* H) by rewrite im_fH im_fK.
+have eq_fHK: {in H :&: K, fH =1 fK}.
+ by move=> z; rewrite ziHK => Zz; rewrite fHZ ?fKZ /= ?eqZHK // def_gH def_gK.
+exists (cprodm_morphism defG cfHK eq_fHK).
+rewrite injm_cprodm injfH injfK im_cprodm im_fH im_fK defHK.
+rewrite -morphimIdom ziHK -eqZHK injm_center // im_fH eqxx.
+split=> //= z; rewrite {1}defZ => Zz; have [Hz _] := setIP Zz.
+by rewrite cprodmEl // fHZ // def_gH.
+Qed.
+
+End Product.
+
+Section CprodBy.
+
+Variables gTH gTK : finGroupType.
+Variables (H : {group gTH}) (K : {group gTK}) (gz : {morphism 'Z(H) >-> gTK}).
+
+Definition ker_cprod_by of isom 'Z(H) 'Z(K) gz :=
+ [set xy | let: (x, y) := xy in (x \in 'Z(H)) && (y == (gz x)^-1)].
+
+Hypothesis isoZ : isom 'Z(H) 'Z(K) gz.
+Let kerHK := ker_cprod_by isoZ.
+
+Let injgz : 'injm gz. Proof. by case/isomP: isoZ. Qed.
+Let gzZ : gz @* 'Z(H) = 'Z(K). Proof. by case/isomP: isoZ. Qed.
+Let gzZchar : gz @* 'Z(H) \char 'Z(K). Proof. by rewrite gzZ char_refl. Qed.
+Let sgzZZ : gz @* 'Z(H) \subset 'Z(K) := char_sub gzZchar.
+Let sZH := center_sub H.
+Let sZK := center_sub K.
+Let sgzZG : gz @* 'Z(H) \subset K := subset_trans sgzZZ sZK.
+
+Lemma ker_cprod_by_is_group : group_set kerHK.
+Proof.
+apply/group_setP; rewrite inE /= group1 morph1 invg1 /=.
+split=> // [[x1 y1] [x2 y2]].
+rewrite inE /= => /andP[Zx1 /eqP->]; have [_ cGx1] := setIP Zx1.
+rewrite inE /= => /andP[Zx2 /eqP->]; have [Gx2 _] := setIP Zx2.
+by rewrite inE /= groupM //= -invMg (centP cGx1) // morphM.
+Qed.
+Canonical ker_cprod_by_group := Group ker_cprod_by_is_group.
+
+Lemma ker_cprod_by_central : kerHK \subset 'Z(setX H K).
+Proof.
+rewrite -(center_dprod (setX_dprod H K)) -morphim_pairg1 -morphim_pair1g.
+rewrite -!injm_center ?subsetT ?injm_pair1g ?injm_pairg1 //=.
+rewrite morphim_pairg1 morphim_pair1g setX_dprod.
+apply/subsetP=> [[x y]]; rewrite inE => /andP[Zx /eqP->].
+by rewrite inE /= Zx groupV (subsetP sgzZZ) ?mem_morphim.
+Qed.
+
+Fact cprod_by_key : unit. Proof. by []. Qed.
+Definition cprod_by_def := subFinGroupType [group of setX H K / kerHK].
+Definition cprod_by := locked_with cprod_by_key cprod_by_def.
+Local Notation C := [set: FinGroup.arg_sort (FinGroup.base cprod_by)].
+
+Definition in_cprod : gTH * gTK -> cprod_by :=
+ let: tt as k := cprod_by_key return _ -> locked_with k cprod_by_def in
+ subg _ \o coset kerHK.
+
+Lemma in_cprodM : {in setX H K &, {morph in_cprod : u v / u * v}}.
+Proof.
+rewrite /in_cprod /cprod_by; case: cprod_by_key => /= u v Gu Gv.
+have nkerHKG := normal_norm (sub_center_normal ker_cprod_by_central).
+by rewrite -!morphM ?mem_quotient // (subsetP nkerHKG).
+Qed.
+Canonical in_cprod_morphism := Morphism in_cprodM.
+
+Lemma ker_in_cprod : 'ker in_cprod = kerHK.
+Proof.
+transitivity ('ker (subg [group of setX H K / kerHK] \o coset kerHK)).
+ rewrite /ker /morphpre /= /in_cprod /cprod_by; case: cprod_by_key => /=.
+ by rewrite ['N(_) :&: _]quotientGK ?sub_center_normal ?ker_cprod_by_central.
+by rewrite ker_comp ker_subg -kerE ker_coset.
+Qed.
+
+Lemma cpairg1_dom : H \subset 'dom (in_cprod \o @pairg1 gTH gTK).
+Proof. by rewrite -sub_morphim_pre ?subsetT // morphim_pairg1 setXS ?sub1G. Qed.
+
+Lemma cpair1g_dom : K \subset 'dom (in_cprod \o @pair1g gTH gTK).
+Proof. by rewrite -sub_morphim_pre ?subsetT // morphim_pair1g setXS ?sub1G. Qed.
+
+Definition cpairg1 := tag (restrmP _ cpairg1_dom).
+Definition cpair1g := tag (restrmP _ cpair1g_dom).
+
+Local Notation CH := (mfun cpairg1 @* gval H).
+Local Notation CK := (mfun cpair1g @* gval K).
+
+Lemma injm_cpairg1 : 'injm cpairg1.
+Proof.
+rewrite /cpairg1; case: restrmP => _ [_ -> _ _].
+rewrite ker_comp ker_in_cprod; apply/subsetP=> x; rewrite 5!inE /=.
+by case/and3P=> _ Zx; rewrite inE eq_sym (inv_eq invgK) invg1 morph_injm_eq1.
+Qed.
+Let injH := injm_cpairg1.
+
+Lemma injm_cpair1g : 'injm cpair1g.
+Proof.
+rewrite /cpair1g; case: restrmP => _ [_ -> _ _].
+rewrite ker_comp ker_in_cprod; apply/subsetP=> y; rewrite !inE /= morph1 invg1.
+by case/and3P.
+Qed.
+Let injK := injm_cpair1g.
+
+Lemma im_cpair_cent : CK \subset 'C(CH).
+Proof.
+rewrite /cpairg1 /cpair1g; do 2!case: restrmP => _ [_ _ _ -> //].
+rewrite !morphim_comp morphim_cents // morphim_pair1g morphim_pairg1.
+by case/dprodP: (setX_dprod H K).
+Qed.
+Hint Resolve im_cpair_cent.
+
+Lemma im_cpair : CH * CK = C.
+Proof.
+rewrite /cpairg1 /cpair1g; do 2!case: restrmP => _ [_ _ _ -> //].
+rewrite !morphim_comp -morphimMl morphim_pairg1 ?setXS ?sub1G //.
+rewrite morphim_pair1g setX_prod morphimEdom /= /in_cprod /cprod_by.
+by case: cprod_by_key; rewrite /= imset_comp imset_coset -morphimEdom im_subg.
+Qed.
+
+Lemma im_cpair_cprod : CH \* CK = C. Proof. by rewrite cprodE ?im_cpair. Qed.
+
+Lemma eq_cpairZ : {in 'Z(H), cpairg1 =1 cpair1g \o gz}.
+Proof.
+rewrite /cpairg1 /cpair1g => z1 Zz1; set z2 := gz z1.
+have Zz2: z2 \in 'Z(K) by rewrite (subsetP sgzZZ) ?mem_morphim.
+have [[Gz1 _] [/= Gz2 _]]:= (setIP Zz1, setIP Zz2).
+do 2![case: restrmP => f /= [df _ _ _]; rewrite {f}df].
+apply/rcoset_kerP; rewrite ?inE ?group1 ?andbT //.
+by rewrite ker_in_cprod mem_rcoset inE /= invg1 mulg1 mul1g Zz1 /=.
+Qed.
+
+Lemma setI_im_cpair : CH :&: CK = 'Z(CH).
+Proof.
+apply/eqP; rewrite eqEsubset setIS //=.
+rewrite subsetI center_sub -injm_center //.
+rewrite (eq_in_morphim _ eq_cpairZ); first by rewrite morphim_comp morphimS.
+by rewrite !(setIidPr _) // -sub_morphim_pre.
+Qed.
+
+Lemma cpair1g_center : cpair1g @* 'Z(K) = 'Z(C).
+Proof.
+case/cprodP: (center_cprod im_cpair_cprod) => _ <- _.
+by rewrite injm_center // -setI_im_cpair mulSGid //= setIC setIS 1?centsC.
+Qed.
+
+(* Uses gzZ. *)
+Lemma cpair_center_id : 'Z(CH) = 'Z(CK).
+Proof.
+rewrite -!injm_center // -gzZ -morphim_comp; apply: eq_in_morphim eq_cpairZ.
+by rewrite !(setIidPr _) // -sub_morphim_pre.
+Qed.
+
+(* Uses gzZ. *)
+Lemma cpairg1_center : cpairg1 @* 'Z(H) = 'Z(C).
+Proof. by rewrite -cpair1g_center !injm_center // cpair_center_id. Qed.
+
+Section ExtCprodm.
+
+Variable rT : finGroupType.
+Variables (fH : {morphism H >-> rT}) (fK : {morphism K >-> rT}).
+Hypothesis cfHK : fK @* K \subset 'C(fH @* H).
+Hypothesis eq_fHK : {in 'Z(H), fH =1 fK \o gz}.
+
+Let gH := ifactm fH injm_cpairg1.
+Let gK := ifactm fK injm_cpair1g.
+
+Lemma xcprodm_cent : gK @* CK \subset 'C(gH @* CH).
+Proof. by rewrite !im_ifactm. Qed.
+
+Lemma xcprodmI : {in CH :&: CK, gH =1 gK}.
+Proof.
+rewrite setI_im_cpair -injm_center // => fHx; case/morphimP=> x Gx Zx ->{fHx}.
+by rewrite {2}eq_cpairZ //= ?ifactmE ?eq_fHK //= (subsetP sgzZG) ?mem_morphim.
+Qed.
+
+Definition xcprodm := cprodm im_cpair_cprod xcprodm_cent xcprodmI.
+Canonical xcprod_morphism := [morphism of xcprodm].
+
+Lemma xcprodmEl : {in H, forall x, xcprodm (cpairg1 x) = fH x}.
+Proof. by move=> x Hx; rewrite /xcprodm cprodmEl ?mem_morphim ?ifactmE. Qed.
+
+Lemma xcprodmEr : {in K, forall y, xcprodm (cpair1g y) = fK y}.
+Proof. by move=> y Ky; rewrite /xcprodm cprodmEr ?mem_morphim ?ifactmE. Qed.
+
+Lemma xcprodmE :
+ {in H & K, forall x y, xcprodm (cpairg1 x * cpair1g y) = fH x * fK y}.
+Proof.
+by move=> x y Hx Ky; rewrite /xcprodm cprodmE ?mem_morphim ?ifactmE.
+Qed.
+
+Lemma im_xcprodm : xcprodm @* C = fH @* H * fK @* K.
+Proof. by rewrite -im_cpair morphim_cprodm // !im_ifactm. Qed.
+
+Lemma im_xcprodml A : xcprodm @* (cpairg1 @* A) = fH @* A.
+Proof.
+rewrite -!(morphimIdom _ A) morphim_cprodml ?morphimS ?subsetIl //.
+by rewrite morphim_ifactm ?subsetIl.
+Qed.
+
+Lemma im_xcprodmr A : xcprodm @* (cpair1g @* A) = fK @* A.
+Proof.
+rewrite -!(morphimIdom _ A) morphim_cprodmr ?morphimS ?subsetIl //.
+by rewrite morphim_ifactm ?subsetIl.
+Qed.
+
+Lemma injm_xcprodm : 'injm xcprodm = 'injm fH && 'injm fK.
+Proof.
+rewrite injm_cprodm !ker_ifactm !subG1 !morphim_injm_eq1 ?subsetIl // -!subG1.
+apply: andb_id2l => /= injfH; apply: andb_idr => _.
+rewrite !im_ifactm // -(morphimIdom gH) setI_im_cpair -injm_center //.
+rewrite morphim_ifactm // eqEsubset subsetI morphimS //=.
+rewrite {1}injm_center // setIS //=.
+rewrite (eq_in_morphim _ eq_fHK); first by rewrite morphim_comp morphimS.
+by rewrite !(setIidPr _) // -sub_morphim_pre.
+Qed.
+
+End ExtCprodm.
+
+(* Uses gzZchar. *)
+Lemma Aut_cprod_by_full :
+ Aut_in (Aut H) 'Z(H) \isog Aut 'Z(H) ->
+ Aut_in (Aut K) 'Z(K) \isog Aut 'Z(K) ->
+ Aut_in (Aut C) 'Z(C) \isog Aut 'Z(C).
+Proof.
+move=> AutZinH AutZinK.
+have Cfull:= Aut_cprod_full im_cpair_cprod cpair_center_id.
+by rewrite Cfull // -injm_center // injm_Aut_full ?center_sub.
+Qed.
+
+Section Isomorphism.
+
+Let gzZ_lone (Y : {group gTK}) :
+ Y \subset 'Z(K) -> gz @* 'Z(H) \isog Y -> gz @* 'Z(H) = Y.
+Proof.
+move=> sYZ isoY; apply/eqP.
+by rewrite eq_sym eqEcard (card_isog isoY) gzZ sYZ /=.
+Qed.
+
+Variables (rT : finGroupType) (GH GK G : {group rT}).
+Hypotheses (defG : GH \* GK = G) (ziGHK : GH :&: GK = 'Z(GH)).
+Hypothesis AutZHfull : Aut_in (Aut H) 'Z(H) \isog Aut 'Z(H).
+Hypotheses (isoGH : GH \isog H) (isoGK : GK \isog K).
+
+(* Uses gzZ_lone *)
+Lemma cprod_by_uniq :
+ exists f : {morphism G >-> cprod_by},
+ [/\ isom G C f, f @* GH = CH & f @* GK = CK].
+Proof.
+have [_ defGHK cGKH] := cprodP defG.
+have AutZinH := Aut_sub_fullP sZH AutZHfull.
+have [fH injfH defGH]:= isogP (isog_symr isoGH).
+have [fK injfK defGK]:= isogP (isog_symr isoGK).
+have sfHZfK: fH @* 'Z(H) \subset fK @* K.
+ by rewrite injm_center //= defGH defGK -ziGHK subsetIr.
+have gzZ_id: gz @* 'Z(H) = invm injfK @* (fH @* 'Z(H)).
+ apply: gzZ_lone => /=.
+ rewrite injm_center // defGH -ziGHK sub_morphim_pre /= ?defGK ?subsetIr //.
+ by rewrite setIC morphpre_invm injm_center // defGK setIS 1?centsC.
+ rewrite -morphim_comp.
+ apply: isog_trans (sub_isog _ _); first by rewrite isog_sym sub_isog.
+ by rewrite -sub_morphim_pre.
+ by rewrite !injm_comp ?injm_invm.
+have: 'dom (invm injfH \o fK \o gz) = 'Z(H).
+ rewrite /dom /= -(morphpreIdom gz); apply/setIidPl.
+ by rewrite -2?sub_morphim_pre // gzZ_id morphim_invmE morphpreK ?morphimS.
+case/domP=> gzH [def_gzH ker_gzH _ im_gzH].
+have{ker_gzH} injgzH: 'injm gzH by rewrite ker_gzH !injm_comp ?injm_invm.
+have{AutZinH} [|gH [injgH gH_H def_gH]] := AutZinH _ injgzH.
+ by rewrite im_gzH !morphim_comp /= gzZ_id !morphim_invmE morphpreK ?injmK.
+have: 'dom (fH \o gH) = H by rewrite /dom /= -{3}gH_H injmK.
+case/domP=> gfH [def_gfH ker_gfH _ im_gfH].
+have{im_gfH} gfH_H: gfH @* H = GH by rewrite im_gfH morphim_comp gH_H.
+have cgfHfK: fK @* K \subset 'C(gfH @* H) by rewrite gfH_H defGK.
+have eq_gfHK: {in 'Z(H), gfH =1 fK \o gz}.
+ move=> z Zz; rewrite def_gfH /= def_gH //= def_gzH /= invmK //.
+ have {Zz}: gz z \in gz @* 'Z(H) by rewrite mem_morphim.
+ rewrite gzZ_id morphim_invmE; case/morphpreP=> _.
+ exact: (subsetP (morphimS _ _)).
+pose f := xcprodm cgfHfK eq_gfHK.
+have injf: 'injm f by rewrite injm_xcprodm ker_gfH injm_comp.
+have fCH: f @* CH = GH by rewrite im_xcprodml gfH_H.
+have fCK: f @* CK = GK by rewrite im_xcprodmr defGK.
+have fC: f @* C = G by rewrite im_xcprodm gfH_H defGK defGHK.
+have [f' [_ ker_f' _ im_f']] := domP (invm_morphism injf) fC.
+exists f'; rewrite -fCH -fCK !{1}im_f' !{1}morphim_invm ?subsetT //.
+by split=> //; apply/isomP; rewrite ker_f' injm_invm im_f' -fC im_invm.
+Qed.
+
+Lemma isog_cprod_by : G \isog C.
+Proof. by have [f [isoG _ _]] := cprod_by_uniq; exact: isom_isog isoG. Qed.
+
+End Isomorphism.
+
+End CprodBy.
+
+Section ExtCprod.
+Import finfun.
+
+Variables gTH gTK : finGroupType.
+Variables (H : {group gTH}) (K : {group gTK}).
+
+Let gt_ b := if b then gTK else gTH.
+Local Notation isob := ('Z(H) \isog 'Z(K)) (only parsing).
+Let G_ b := if b as b' return {group gt_ b'} then K else H.
+
+Lemma xcprod_subproof :
+ {gz : {morphism 'Z(H) >-> gt_ isob} | isom 'Z(H) 'Z(G_ isob) gz}.
+Proof.
+case: (pickP [pred f : {ffun _} | misom 'Z(H) 'Z(K) f]) => [f isoZ | no_f].
+ rewrite (misom_isog isoZ); case/andP: isoZ => fM isoZ.
+ by exists [morphism of morphm fM].
+move/pred0P: no_f => not_isoZ; rewrite [isob](congr1 negb not_isoZ).
+by exists (idm_morphism _); apply/isomP; rewrite injm_idm im_idm.
+Qed.
+
+Definition xcprod := cprod_by (svalP xcprod_subproof).
+
+Inductive xcprod_spec : finGroupType -> Prop :=
+ XcprodSpec gz isoZ : xcprod_spec (@cprod_by gTH gTK H K gz isoZ).
+
+Lemma xcprodP : 'Z(H) \isog 'Z(K) -> xcprod_spec xcprod.
+Proof. by rewrite /xcprod => isoZ; move: xcprod_subproof; rewrite isoZ. Qed.
+
+Lemma isog_xcprod (rT : finGroupType) (GH GK G : {group rT}) :
+ Aut_in (Aut H) 'Z(H) \isog Aut 'Z(H) ->
+ GH \isog H -> GK \isog K -> GH \* GK = G -> 'Z(GH) = 'Z(GK) ->
+ G \isog [set: xcprod].
+Proof.
+move=> AutZinH isoGH isoGK defG eqZGHK; have [_ _ cGHK] := cprodP defG.
+have [|gz isoZ] := xcprodP.
+ have [[fH injfH <-] [fK injfK <-]] := (isogP isoGH, isogP isoGK).
+ rewrite -!injm_center -?(isog_transl _ (sub_isog _ _)) ?center_sub //=.
+ by rewrite eqZGHK sub_isog ?center_sub.
+rewrite (isog_cprod_by _ defG) //.
+by apply/eqP; rewrite eqEsubset setIS // subsetI {2}eqZGHK !center_sub.
+Qed.
+
+End ExtCprod.
+
+Section IterCprod.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Fixpoint ncprod_def n : finGroupType :=
+ if n is n'.+1 then xcprod G [set: ncprod_def n']
+ else [finGroupType of subg_of 'Z(G)].
+Fact ncprod_key : unit. Proof. by []. Qed.
+Definition ncprod := locked_with ncprod_key ncprod_def.
+
+Local Notation G_ n := [set: gsort (ncprod n)].
+
+Lemma ncprod0 : G_ 0 \isog 'Z(G).
+Proof. by rewrite [ncprod]unlock isog_sym isog_subg. Qed.
+
+Lemma center_ncprod0 : 'Z(G_ 0) = G_ 0.
+Proof. by apply: center_idP; rewrite (isog_abelian ncprod0) center_abelian. Qed.
+
+Lemma center_ncprod n : 'Z(G_ n) \isog 'Z(G).
+Proof.
+elim: n => [|n]; first by rewrite center_ncprod0 ncprod0.
+rewrite [ncprod]unlock=> /isog_symr/xcprodP[gz isoZ] /=.
+by rewrite -cpairg1_center isog_sym sub_isog ?center_sub ?injm_cpairg1.
+Qed.
+
+Lemma ncprodS n : xcprod_spec G [set: ncprod n] (ncprod n.+1).
+Proof.
+by have:= xcprodP (isog_symr (center_ncprod n)); rewrite [ncprod]unlock.
+Qed.
+
+Lemma ncprod1 : G_ 1 \isog G.
+Proof.
+case: ncprodS => gz isoZ; rewrite isog_sym /= -im_cpair.
+rewrite mulGSid /=; first by rewrite sub_isog ?injm_cpairg1.
+rewrite -{3}center_ncprod0 injm_center ?injm_cpair1g //.
+by rewrite -cpair_center_id center_sub.
+Qed.
+
+Lemma Aut_ncprod_full n :
+ Aut_in (Aut G) 'Z(G) \isog Aut 'Z(G) ->
+ Aut_in (Aut (G_ n)) 'Z(G_ n) \isog Aut 'Z(G_ n).
+Proof.
+move=> AutZinG; elim: n => [|n IHn].
+ by rewrite center_ncprod0; apply/Aut_sub_fullP=> // g injg gG0; exists g.
+by case: ncprodS => gz isoZ; exact: Aut_cprod_by_full.
+Qed.
+
+End IterCprod.
+
+
diff --git a/mathcomp/solvable/commutator.v b/mathcomp/solvable/commutator.v
new file mode 100644
index 0000000..81be11d
--- /dev/null
+++ b/mathcomp/solvable/commutator.v
@@ -0,0 +1,362 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat fintype bigop finset.
+Require Import binomial fingroup morphism automorphism quotient gfunctor.
+
+(******************************************************************************)
+(* This files contains the proofs of several key properties of commutators, *)
+(* including the Hall-Witt identity and the Three Subgroup Lemma. *)
+(* The definition and notation for both pointwise and set wise commutators *)
+(* ([~x, y, ...] and [~: A, B ,...], respectively) are given in fingroup.v *)
+(* This file defines the derived group series: *)
+(* G^`(0) == G *)
+(* G^`(n.+1) == [~: G^`(n), G^`(n)] *)
+(* as several classical results involve the (first) derived group G^`(1), *)
+(* such as the equivalence H <| G /\ G / H abelian <-> G^`(1) \subset H. The *)
+(* connection between the derived series and solvable groups will only be *)
+(* established in nilpotent.v, however. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Definition derived_at_rec n (gT : finGroupType) (A : {set gT}) :=
+ iter n (fun B => [~: B, B]) A.
+
+(* Note: 'nosimpl' MUST be used outside of a section -- the end of section *)
+(* "cooking" destroys it. *)
+Definition derived_at := nosimpl derived_at_rec.
+
+Arguments Scope derived_at [nat_scope _ group_scope].
+Notation "G ^` ( n )" := (derived_at n G) : group_scope.
+
+Section DerivedBasics.
+
+Variables gT : finGroupType.
+Implicit Type A : {set gT}.
+Implicit Types G : {group gT}.
+
+Lemma derg0 A : A^`(0) = A. Proof. by []. Qed.
+Lemma derg1 A : A^`(1) = [~: A, A]. Proof. by []. Qed.
+Lemma dergSn n A : A^`(n.+1) = [~: A^`(n), A^`(n)]. Proof. by []. Qed.
+
+Lemma der_group_set G n : group_set G^`(n).
+Proof. by case: n => [|n]; exact: groupP. Qed.
+
+Canonical derived_at_group G n := Group (der_group_set G n).
+
+End DerivedBasics.
+
+Notation "G ^` ( n )" := (derived_at_group G n) : Group_scope.
+
+Section Basic_commutator_properties.
+
+Variable gT : finGroupType.
+Implicit Types x y z : gT.
+
+Lemma conjg_mulR x y : x ^ y = x * [~ x, y].
+Proof. by rewrite mulKVg. Qed.
+
+Lemma conjg_Rmul x y : x ^ y = [~ y, x^-1] * x.
+Proof. by rewrite commgEr invgK mulgKV. Qed.
+
+Lemma commMgJ x y z : [~ x * y, z] = [~ x, z] ^ y * [~ y, z].
+Proof. by rewrite !commgEr conjgM mulgA -conjMg mulgK. Qed.
+
+Lemma commgMJ x y z : [~ x, y * z] = [~ x, z] * [~ x, y] ^ z.
+Proof. by rewrite !commgEl conjgM -mulgA -conjMg mulKVg. Qed.
+
+Lemma commMgR x y z : [~ x * y, z] = [~ x, z] * [~ x, z, y] * [~ y, z].
+Proof. by rewrite commMgJ conjg_mulR. Qed.
+
+Lemma commgMR x y z : [~ x, y * z] = [~ x, z] * [~ x, y] * [~ x, y, z].
+Proof. by rewrite commgMJ conjg_mulR mulgA. Qed.
+
+Lemma Hall_Witt_identity x y z :
+ [~ x, y^-1, z] ^ y * [~ y, z^-1, x] ^ z * [~ z, x^-1, y] ^ x = 1.
+Proof. (* gsimpl *)
+pose a x y z : gT := x * z * y ^ x.
+suffices{x y z} hw_aux x y z: [~ x, y^-1, z] ^ y = (a x y z)^-1 * (a y z x).
+ by rewrite !hw_aux 2!mulgA !mulgK mulVg.
+by rewrite commgEr conjMg -conjgM -conjg_Rmul 2!invMg conjgE !mulgA.
+Qed.
+
+(* the following properties are useful for studying p-groups of class 2 *)
+
+Section LeftComm.
+
+Variables (i : nat) (x y : gT).
+Hypothesis cxz : commute x [~ x, y].
+
+Lemma commVg : [~ x^-1, y] = [~ x, y]^-1.
+Proof.
+apply/eqP; rewrite commgEl eq_sym eq_invg_mul invgK mulgA -cxz.
+by rewrite -conjg_mulR -conjMg mulgV conj1g.
+Qed.
+
+Lemma commXg : [~ x ^+ i, y] = [~ x, y] ^+ i.
+Proof.
+elim: i => [|i' IHi]; first exact: comm1g.
+by rewrite !expgS commMgJ /conjg commuteX // mulKg IHi.
+Qed.
+
+End LeftComm.
+
+Section RightComm.
+
+Variables (i : nat) (x y : gT).
+Hypothesis cyz : commute y [~ x, y].
+Let cyz' := commuteV cyz.
+
+Lemma commgV : [~ x, y^-1] = [~ x, y]^-1.
+Proof. by rewrite -invg_comm commVg -(invg_comm x y) ?invgK. Qed.
+
+Lemma commgX : [~ x, y ^+ i] = [~ x, y] ^+ i.
+Proof. by rewrite -invg_comm commXg -(invg_comm x y) ?expgVn ?invgK. Qed.
+
+End RightComm.
+
+Section LeftRightComm.
+
+Variables (i j : nat) (x y : gT).
+Hypotheses (cxz : commute x [~ x, y]) (cyz : commute y [~ x, y]).
+
+Lemma commXXg : [~ x ^+ i, y ^+ j] = [~ x, y] ^+ (i * j).
+Proof. rewrite expgM commgX commXg //; exact: commuteX. Qed.
+
+Lemma expMg_Rmul : (y * x) ^+ i = y ^+ i * x ^+ i * [~ x, y] ^+ 'C(i, 2).
+Proof.
+rewrite -triangular_sum; symmetry.
+elim: i => [|k IHk] /=; first by rewrite big_geq ?mulg1.
+rewrite big_nat_recr //= addnC expgD !expgS -{}IHk !mulgA; congr (_ * _).
+by rewrite -!mulgA commuteX2 // -commgX // [mulg y]lock 3!mulgA -commgC.
+Qed.
+
+End LeftRightComm.
+
+End Basic_commutator_properties.
+
+(***** Set theoretic commutators *****)
+Section Commutator_properties.
+
+Variable gT : finGroupType.
+Implicit Type (rT : finGroupType) (A B C : {set gT}) (D G H K : {group gT}).
+
+Lemma commG1 A : [~: A, 1] = 1.
+Proof. by apply/commG1P; rewrite centsC sub1G. Qed.
+
+Lemma comm1G A : [~: 1, A] = 1.
+Proof. by rewrite commGC commG1. Qed.
+
+Lemma commg_sub A B : [~: A, B] \subset A <*> B.
+Proof. by rewrite comm_subG // (joing_subl, joing_subr). Qed.
+
+Lemma commg_norml G A : G \subset 'N([~: G, A]).
+Proof.
+apply/subsetP=> x Gx; rewrite inE -genJ gen_subG.
+apply/subsetP=> _ /imsetP[_ /imset2P[y z Gy Az ->] ->].
+by rewrite -(mulgK [~ x, z] (_ ^ x)) -commMgJ !(mem_commg, groupMl, groupV).
+Qed.
+
+Lemma commg_normr G A : G \subset 'N([~: A, G]).
+Proof. by rewrite commGC commg_norml. Qed.
+
+Lemma commg_norm G H : G <*> H \subset 'N([~: G, H]).
+Proof. by rewrite join_subG ?commg_norml ?commg_normr. Qed.
+
+Lemma commg_normal G H : [~: G, H] <| G <*> H.
+Proof. by rewrite /(_ <| _) commg_sub commg_norm. Qed.
+
+Lemma normsRl A G B : A \subset G -> A \subset 'N([~: G, B]).
+Proof. by move=> sAG; exact: subset_trans (commg_norml G B). Qed.
+
+Lemma normsRr A G B : A \subset G -> A \subset 'N([~: B, G]).
+Proof. by move=> sAG; exact: subset_trans (commg_normr G B). Qed.
+
+Lemma commg_subr G H : ([~: G, H] \subset H) = (G \subset 'N(H)).
+Proof.
+rewrite gen_subG; apply/subsetP/subsetP=> [sRH x Gx | nGH xy].
+ rewrite inE; apply/subsetP=> _ /imsetP[y Ky ->].
+ by rewrite conjg_Rmul groupMr // sRH // mem_imset2 ?groupV.
+case/imset2P=> x y Gx Hy ->{xy}.
+by rewrite commgEr groupMr // memJ_norm (groupV, nGH).
+Qed.
+
+Lemma commg_subl G H : ([~: G, H] \subset G) = (H \subset 'N(G)).
+Proof. by rewrite commGC commg_subr. Qed.
+
+Lemma commg_subI A B G H :
+ A \subset 'N_G(H) -> B \subset 'N_H(G) -> [~: A, B] \subset G :&: H.
+Proof.
+rewrite !subsetI -(gen_subG _ 'N(G)) -(gen_subG _ 'N(H)).
+rewrite -commg_subr -commg_subl; case/andP=> sAG sRH; case/andP=> sBH sRG.
+by rewrite (subset_trans _ sRG) ?(subset_trans _ sRH) ?commgSS ?subset_gen.
+Qed.
+
+Lemma quotient_cents2 A B K :
+ A \subset 'N(K) -> B \subset 'N(K) ->
+ (A / K \subset 'C(B / K)) = ([~: A, B] \subset K).
+Proof.
+move=> nKA nKB.
+by rewrite (sameP commG1P trivgP) /= -quotientR // quotient_sub1 // comm_subG.
+Qed.
+
+Lemma quotient_cents2r A B K :
+ [~: A, B] \subset K -> (A / K) \subset 'C(B / K).
+Proof.
+move=> sABK; rewrite -2![_ / _]morphimIdom -!quotientE.
+by rewrite quotient_cents2 ?subsetIl ?(subset_trans _ sABK) ?commgSS ?subsetIr.
+Qed.
+
+Lemma sub_der1_norm G H : G^`(1) \subset H -> H \subset G -> G \subset 'N(H).
+Proof.
+by move=> sG'H sHG; rewrite -commg_subr (subset_trans _ sG'H) ?commgS.
+Qed.
+
+Lemma sub_der1_normal G H : G^`(1) \subset H -> H \subset G -> H <| G.
+Proof. by move=> sG'H sHG; rewrite /(H <| G) sHG sub_der1_norm. Qed.
+
+Lemma sub_der1_abelian G H : G^`(1) \subset H -> abelian (G / H).
+Proof. by move=> sG'H; exact: quotient_cents2r. Qed.
+
+Lemma der1_min G H : G \subset 'N(H) -> abelian (G / H) -> G^`(1) \subset H.
+Proof. by move=> nHG abGH; rewrite -quotient_cents2. Qed.
+
+Lemma der_abelian n G : abelian (G^`(n) / G^`(n.+1)).
+Proof. by rewrite sub_der1_abelian // der_subS. Qed.
+
+Lemma commg_normSl G H K : G \subset 'N(H) -> [~: G, H] \subset 'N([~: K, H]).
+Proof. by move=> nHG; rewrite normsRr // commg_subr. Qed.
+
+Lemma commg_normSr G H K : G \subset 'N(H) -> [~: H, G] \subset 'N([~: H, K]).
+Proof. by move=> nHG; rewrite !(commGC H) commg_normSl. Qed.
+
+Lemma commMGr G H K : [~: G, K] * [~: H, K] \subset [~: G * H , K].
+Proof. by rewrite mul_subG ?commSg ?(mulG_subl, mulG_subr). Qed.
+
+Lemma commMG G H K :
+ H \subset 'N([~: G, K]) -> [~: G * H , K] = [~: G, K] * [~: H, K].
+Proof.
+move=> nRH; apply/eqP; rewrite eqEsubset commMGr andbT.
+have nRHK: [~: H, K] \subset 'N([~: G, K]) by rewrite comm_subG ?commg_normr.
+have defM := norm_joinEr nRHK; rewrite -defM gen_subG /=.
+apply/subsetP=> _ /imset2P[_ z /imset2P[x y Gx Hy ->] Kz ->].
+by rewrite commMgJ {}defM mem_mulg ?memJ_norm ?mem_commg // (subsetP nRH).
+Qed.
+
+Lemma comm3G1P A B C :
+ reflect {in A & B & C, forall h k l, [~ h, k, l] = 1} ([~: A, B, C] :==: 1).
+Proof.
+have R_C := sameP trivgP commG1P.
+rewrite -subG1 R_C gen_subG -{}R_C gen_subG.
+apply: (iffP subsetP) => [cABC x y z Ax By Cz | cABC xyz].
+ by apply/set1P; rewrite cABC // !mem_imset2.
+by case/imset2P=> _ z /imset2P[x y Ax By ->] Cz ->; rewrite cABC.
+Qed.
+
+Lemma three_subgroup G H K :
+ [~: G, H, K] :=: 1 -> [~: H, K, G] :=: 1-> [~: K, G, H] :=: 1.
+Proof.
+move/eqP/comm3G1P=> cGHK /eqP/comm3G1P cHKG.
+apply/eqP/comm3G1P=> x y z Kx Gy Hz; symmetry.
+rewrite -(conj1g y) -(Hall_Witt_identity y^-1 z x) invgK.
+by rewrite cGHK ?groupV // cHKG ?groupV // !conj1g !mul1g conjgKV.
+Qed.
+
+Lemma der1_joing_cycles (x y : gT) :
+ let XY := <[x]> <*> <[y]> in let xy := [~ x, y] in
+ xy \in 'C(XY) -> XY^`(1) = <[xy]>.
+Proof.
+rewrite joing_idl joing_idr /= -sub_cent1 => /norms_gen nRxy.
+apply/eqP; rewrite eqEsubset cycle_subG mem_commg ?mem_gen ?set21 ?set22 //.
+rewrite der1_min // quotient_gen -1?gen_subG // quotientU abelian_gen.
+rewrite /abelian subUset centU !subsetI andbC centsC -andbA -!abelianE.
+rewrite !quotient_abelian ?(abelianS (subset_gen _) (cycle_abelian _)) //=.
+by rewrite andbb quotient_cents2r ?genS // /commg_set imset2_set1l imset_set1.
+Qed.
+
+Lemma commgAC G x y z : x \in G -> y \in G -> z \in G ->
+ commute y z -> abelian [~: [set x], G] -> [~ x, y, z] = [~ x, z, y].
+Proof.
+move=> Gx Gy Gz cyz /centsP cRxG; pose cx' u := [~ x^-1, u].
+have xR3 u v: [~ x, u, v] = x^-1 * (cx' u * cx' v) * x ^ (u * v).
+ rewrite mulgA -conjg_mulR conjVg [cx' v]commgEl mulgA -invMg.
+ by rewrite -mulgA conjgM -conjMg -!commgEl.
+suffices RxGcx' u: u \in G -> cx' u \in [~: [set x], G].
+ by rewrite !xR3 {}cyz; congr (_ * _ * _); rewrite cRxG ?RxGcx'.
+move=> Gu; suffices/groupMl <-: [~ x, u] ^ x^-1 \in [~: [set x], G].
+ by rewrite -commMgJ mulgV comm1g group1.
+by rewrite memJ_norm ?mem_commg ?set11 // groupV (subsetP (commg_normr _ _)).
+Qed.
+
+(* Aschbacher, exercise 3.6 (used in proofs of Aschbacher 24.7 and B & G 1.10 *)
+Lemma comm_norm_cent_cent H G K :
+ H \subset 'N(G) -> H \subset 'C(K) -> G \subset 'N(K) ->
+ [~: G, H] \subset 'C(K).
+Proof.
+move=> nGH /centsP cKH nKG; rewrite commGC gen_subG centsC.
+apply/centsP=> x Kx _ /imset2P[y z Hy Gz ->]; red.
+rewrite mulgA -[x * _]cKH ?groupV // -!mulgA; congr (_ * _).
+rewrite (mulgA x) (conjgC x) (conjgCV z) 3!mulgA; congr (_ * _).
+by rewrite -2!mulgA (cKH y) // -mem_conjg (normsP nKG).
+Qed.
+
+Lemma charR H K G : H \char G -> K \char G -> [~: H, K] \char G.
+Proof.
+case/charP=> sHG chH /charP[sKG chK]; apply/charP.
+by split=> [|f infj Gf]; [rewrite comm_subG | rewrite morphimR // chH // chK].
+Qed.
+
+Lemma der_char n G : G^`(n) \char G.
+Proof. by elim: n => [|n IHn]; rewrite ?char_refl // dergSn charR. Qed.
+
+Lemma der_sub n G : G^`(n) \subset G.
+Proof. by rewrite char_sub ?der_char. Qed.
+
+Lemma der_norm n G : G \subset 'N(G^`(n)).
+Proof. by rewrite char_norm ?der_char. Qed.
+
+Lemma der_normal n G : G^`(n) <| G.
+Proof. by rewrite char_normal ?der_char. Qed.
+
+Lemma der_subS n G : G^`(n.+1) \subset G^`(n).
+Proof. by rewrite comm_subG. Qed.
+
+Lemma der_normalS n G : G^`(n.+1) <| G^`(n).
+Proof. by rewrite sub_der1_normal // der_subS. Qed.
+
+Lemma morphim_der rT D (f : {morphism D >-> rT}) n G :
+ G \subset D -> f @* G^`(n) = (f @* G)^`(n).
+Proof.
+move=> sGD; elim: n => // n IHn.
+by rewrite !dergSn -IHn morphimR ?(subset_trans (der_sub n G)).
+Qed.
+
+Lemma dergS n G H : G \subset H -> G^`(n) \subset H^`(n).
+Proof. by move=> sGH; elim: n => // n IHn; exact: commgSS. Qed.
+
+Lemma quotient_der n G H : G \subset 'N(H) -> G^`(n) / H = (G / H)^`(n).
+Proof. exact: morphim_der. Qed.
+
+Lemma derJ G n x : (G :^ x)^`(n) = G^`(n) :^ x.
+Proof. by elim: n => //= n IHn; rewrite !dergSn IHn -conjsRg. Qed.
+
+Lemma derG1P G : reflect (G^`(1) = 1) (abelian G).
+Proof. exact: commG1P. Qed.
+
+End Commutator_properties.
+
+Implicit Arguments derG1P [gT G].
+
+Lemma der_cont n : GFunctor.continuous (derived_at n).
+Proof. by move=> aT rT G f; rewrite morphim_der. Qed.
+
+Canonical der_igFun n := [igFun by der_sub^~ n & der_cont n].
+Canonical der_gFun n := [gFun by der_cont n].
+Canonical der_mgFun n := [mgFun by dergS^~ n].
+
+Lemma isog_der (aT rT : finGroupType) n (G : {group aT}) (H : {group rT}) :
+ G \isog H -> G^`(n) \isog H^`(n).
+Proof. exact: gFisog. Qed.
diff --git a/mathcomp/solvable/extraspecial.v b/mathcomp/solvable/extraspecial.v
new file mode 100644
index 0000000..c38d17a
--- /dev/null
+++ b/mathcomp/solvable/extraspecial.v
@@ -0,0 +1,833 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import bigop finset prime binomial fingroup morphism perm automorphism.
+Require Import presentation quotient action commutator gproduct gfunctor.
+Require Import ssralg finalg zmodp cyclic pgroup center gseries.
+Require Import nilpotent sylow abelian finmodule matrix maximal extremal.
+
+(******************************************************************************)
+(* This file contains the fine structure thorems for extraspecial p-groups. *)
+(* Together with the material in the maximal and extremal libraries, it *)
+(* completes the coverage of Aschbacher, section 23. *)
+(* We define canonical representatives for the group classes that cover the *)
+(* extremal p-groups (non-abelian p-groups with a cyclic maximal subgroup): *)
+(* 'Mod_m == the modular group of order m, for m = p ^ n, p prime and n >= 3. *)
+(* 'D_m == the dihedral group of order m, for m = 2n >= 4. *)
+(* 'Q_m == the generalized quaternion group of order m, for q = 2 ^ n >= 8. *)
+(* 'SD_m == the semi-dihedral group of order m, for m = 2 ^ n >= 16. *)
+(* In each case the notation is defined in the %type, %g and %G scopes, where *)
+(* it denotes a finGroupType, a full gset and the full group for that type. *)
+(* However each notation is only meaningful under the given conditions, in *)
+(* We construct and study the following extraspecial groups: *)
+(* p^{1+2} == if p is prime, an extraspecial group of order p^3 that has *)
+(* exponent p or 4, and p-rank 2: thus p^{1+2} is isomorphic to *)
+(* 'D_8 if p - 2, and NOT isomorphic to 'Mod_(p^3) if p is odd. *)
+(* p^{1+2*n} == the central product of n copies of p^{1+2}, thus of order *)
+(* p^(1+2*n) if p is a prime, and, when n > 0, a representative *)
+(* of the (unique) isomorphism class of extraspecial groups of *)
+(* order p^(1+2*n), of exponent p or 4, and p-rank n+1. *)
+(* 'D^n == an alternative (and preferred) notation for 2^{1+2*n}, which *)
+(* is isomorphic to the central product of n copies od 'D_8. *)
+(* 'D^n*Q == the central product of 'D^n with 'Q_8, thus isomorphic to *)
+(* all extraspecial groups of order 2 ^ (2 * n + 3) that are *)
+(* not isomorphic to 'D^n.+1 (or, equivalently, have 2-rank n). *)
+(* As in extremal.v, these notations are simultaneously defined in the %type, *)
+(* %g and %G scopes -- depending on the syntactic context, they denote either *)
+(* a finGroupType, the set, or the group of all its elements. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GroupScope GRing.Theory.
+
+Reserved Notation "p ^{1+2}" (at level 2, format "p ^{1+2}").
+Reserved Notation "p ^{1+2* n }"
+ (at level 2, n at level 2, format "p ^{1+2* n }").
+Reserved Notation "''D^' n" (at level 8, n at level 2, format "''D^' n").
+Reserved Notation "''D^' n * 'Q'"
+ (at level 8, n at level 2, format "''D^' n * 'Q'").
+
+Module Pextraspecial.
+
+Section Construction.
+
+Variable p : nat.
+
+Definition act ij (k : 'Z_p) := let: (i, j) := ij in (i + k * j, j).
+Lemma actP : is_action [set: 'Z_p] act.
+Proof.
+apply: is_total_action=> [] [i j] => [|k1 k2] /=; first by rewrite mul0r addr0.
+by rewrite mulrDl addrA.
+Qed.
+Canonical action := Action actP.
+
+Lemma gactP : is_groupAction [set: 'Z_p * 'Z_p] action.
+Proof.
+move=> k _ /=; rewrite inE.
+apply/andP; split; first by apply/subsetP=> ij _; rewrite inE.
+apply/morphicP=> /= [[i1 j1] [i2 j2] _ _].
+by rewrite !permE /= mulrDr -addrA (addrCA i2) (addrA i1).
+Qed.
+Definition groupAction := GroupAction gactP.
+
+Fact gtype_key : unit. Proof. by []. Qed.
+Definition gtype := locked_with gtype_key (sdprod_groupType groupAction).
+
+Definition ngtype := ncprod [set: gtype].
+
+End Construction.
+
+Definition ngtypeQ n := xcprod [set: ngtype 2 n] 'Q_8.
+
+End Pextraspecial.
+
+Notation "p ^{1+2}" := (Pextraspecial.gtype p) : type_scope.
+Notation "p ^{1+2}" := [set: gsort p^{1+2}] : group_scope.
+Notation "p ^{1+2}" := [set: gsort p^{1+2}]%G : Group_scope.
+
+Notation "p ^{1+2* n }" := (Pextraspecial.ngtype p n) : type_scope.
+Notation "p ^{1+2* n }" := [set: gsort p^{1+2*n}] : group_scope.
+Notation "p ^{1+2* n }" := [set: gsort p^{1+2*n}]%G : Group_scope.
+
+Notation "''D^' n" := (Pextraspecial.ngtype 2 n) : type_scope.
+Notation "''D^' n" := [set: gsort 'D^n] : group_scope.
+Notation "''D^' n" := [set: gsort 'D^n]%G : Group_scope.
+
+Notation "''D^' n * 'Q'" := (Pextraspecial.ngtypeQ n) : type_scope.
+Notation "''D^' n * 'Q'" := [set: gsort 'D^n*Q] : group_scope.
+Notation "''D^' n * 'Q'" := [set: gsort 'D^n*Q]%G : Group_scope.
+
+Section ExponentPextraspecialTheory.
+
+Variable p : nat.
+Hypothesis p_pr : prime p.
+Let p_gt1 := prime_gt1 p_pr.
+Let p_gt0 := ltnW p_gt1.
+
+Local Notation gtype := Pextraspecial.gtype.
+Local Notation actp := (Pextraspecial.groupAction p).
+
+Lemma card_pX1p2 : #|p^{1+2}| = (p ^ 3)%N.
+Proof.
+rewrite [@gtype _]unlock -(sdprod_card (sdprod_sdpair _)).
+rewrite !card_injm ?injm_sdpair1 ?injm_sdpair2 // !cardsT card_prod card_ord.
+by rewrite -mulnA Zp_cast.
+Qed.
+
+Lemma Grp_pX1p2 :
+ p^{1+2} \isog Grp (x : y : (x ^+ p, y ^+ p, [~ x, y, x], [~ x, y, y])).
+Proof.
+rewrite [@gtype _]unlock ; apply: intro_isoGrp => [|rT H].
+ apply/existsP; pose x := sdpair1 actp (0, 1)%R; pose y := sdpair2 actp 1%R.
+ exists (x, y); rewrite /= !xpair_eqE; set z := [~ x, y]; set G := _ <*> _.
+ have def_z: z = sdpair1 actp (1, 0)%R.
+ rewrite [z]commgEl -sdpair_act ?inE //=.
+ rewrite -morphV -?morphM ?inE //=; congr (sdpair1 _ (_, _)) => /=.
+ by rewrite mulr1 mulKg.
+ by rewrite mulVg.
+ have def_xi i: x ^+ i = sdpair1 _ (0, i%:R)%R.
+ rewrite -morphX ?inE //; congr (sdpair1 _ _).
+ by apply/eqP; rewrite /eq_op /= !morphX ?inE ?expg1n //=.
+ have def_yi i: y ^+ i = sdpair2 _ i%:R.
+ by rewrite -morphX ?inE //.
+ have def_zi i: z ^+ i = sdpair1 _ (i%:R, 0)%R.
+ rewrite def_z -morphX ?inE //; congr (sdpair1 _ _).
+ by apply/eqP; rewrite /eq_op /= !morphX ?inE ?expg1n ?andbT //=.
+ rewrite def_xi def_yi char_Zp ?morph1 //.
+ rewrite def_z -morphR ?inE // !commgEl -sdpair_act ?inE //= mulr0 addr0.
+ rewrite mulVg -[_ * _]/(_ , _) /= !invg1 mulg1 !mul1g mulVg morph1 !andbT.
+ have Gx: x \in G by rewrite -cycle_subG joing_subl.
+ have Gy: y \in G by rewrite -cycle_subG joing_subr.
+ rewrite eqEsubset subsetT -im_sdpair mulG_subG /= -/G; apply/andP; split.
+ apply/subsetP=> u /morphimP[[i j] _ _ def_u].
+ suffices ->: u = z ^+ i * x ^+ j by rewrite groupMl groupX ?groupR.
+ rewrite def_zi def_xi !natr_Zp -morphM ?inE // def_u.
+ by congr (sdpair1 _ (_, _)); rewrite ?mulg1 ?mul1g.
+ apply/subsetP=> v /morphimP[k _ _ def_v].
+ suffices ->: v = y ^+ k by rewrite groupX.
+ by rewrite def_yi natr_Zp.
+case/existsP=> [[x y] /=]; set z := [~ x, y].
+case/eqP=> defH xp yp /eqP/commgP czx /eqP/commgP czy.
+have zp: z ^+ p = 1 by rewrite -commXg // xp comm1g.
+pose f1 (ij : 'Z_p * 'Z_p) := let: (i, j) := ij in z ^+ i * x ^+ j.
+have f1M: {in setT &, {morph f1 : u v / u * v}}.
+ case=> /= [i1 j1] [i2 j2] _ _ /=; rewrite {3 6}Zp_cast // !expg_mod //.
+ rewrite !expgD !mulgA; congr (_ * _); rewrite -!mulgA; congr (_ * _).
+ by apply: commuteX2.
+pose f2 (k : 'Z_p) := y ^+ k.
+have f2M: {in setT &, {morph f2 : u v / u * v}}.
+ by move=> k1 k2 _ _; rewrite /f2 /= {3}Zp_cast // expg_mod // expgD.
+have actf: {in setT & setT, morph_act actp 'J (Morphism f1M) (Morphism f2M)}.
+ case=> /= i j k _ _; rewrite modnDmr {4}Zp_cast // expg_mod // expgD.
+ rewrite /f2 conjMg {1}/conjg (commuteX2 i k czy) mulKg -mulgA.
+ congr (_ * _); rewrite (commuteX2 _ _ czx) mulnC expgM.
+ by rewrite -commXg // -commgX ?mulKVg // commXg // /commute commuteX.
+apply/homgP; exists (xsdprod_morphism actf).
+apply/eqP; rewrite eqEsubset -{2}defH -genM_join gen_subG /= im_xsdprodm.
+have Hx: x \in H by rewrite -cycle_subG -defH joing_subl.
+have Hy: y \in H by rewrite -cycle_subG -defH joing_subr.
+rewrite mulG_subG -andbA; apply/and3P; split.
+- apply/subsetP=> _ /morphimP[[i j] _ _ -> /=].
+ by rewrite groupMl groupX ?groupR.
+- by apply/subsetP=> _ /morphimP[k _ _ ->]; rewrite groupX.
+rewrite mulgSS ?cycle_subG //= morphimEdom; apply/imsetP.
+ by exists (0, 1)%R; rewrite ?inE //= mul1g.
+by exists 1%R; rewrite ?inE.
+Qed.
+
+Lemma pX1p2_pgroup : p.-group p^{1+2}.
+Proof. by rewrite /pgroup card_pX1p2 pnat_exp pnat_id. Qed.
+
+(* This is part of the existence half of Aschbacher ex. (8.7)(1) *)
+Lemma pX1p2_extraspecial : extraspecial p^{1+2}.
+Proof.
+apply: (p3group_extraspecial pX1p2_pgroup); last first.
+ by rewrite card_pX1p2 pfactorK.
+case/existsP: (isoGrp_hom Grp_pX1p2) card_pX1p2 => [[x y]] /=.
+case/eqP=> <- xp yp _ _ oXY.
+apply: contraL (dvdn_cardMg <[x]> <[y]>) => cXY_XY.
+rewrite -cent_joinEl ?(sub_abelian_cent2 cXY_XY) ?joing_subl ?joing_subr //.
+rewrite oXY -!orderE pfactor_dvdn ?muln_gt0 ?order_gt0 // -leqNgt.
+rewrite -(pfactorK 2 p_pr) dvdn_leq_log ?expn_gt0 ?p_gt0 //.
+by rewrite dvdn_mul ?order_dvdn ?xp ?yp.
+Qed.
+
+(* This is part of the existence half of Aschbacher ex. (8.7)(1) *)
+Lemma exponent_pX1p2 : odd p -> exponent p^{1+2} %| p.
+Proof.
+move=> p_odd; have pG := pX1p2_pgroup.
+have ->: p^{1+2} = 'Ohm_1(p^{1+2}).
+ apply/eqP; rewrite eqEsubset Ohm_sub andbT (OhmE 1 pG).
+ case/existsP: (isoGrp_hom Grp_pX1p2) => [[x y]] /=.
+ case/eqP=> <- xp yp _ _; rewrite joing_idl joing_idr genS //.
+ by rewrite subsetI subset_gen subUset !sub1set !inE xp yp!eqxx.
+rewrite exponent_Ohm1_class2 ?card_pX1p2 ?odd_exp // nil_class2.
+by have [[_ ->] _ ] := pX1p2_extraspecial.
+Qed.
+
+(* This is the uniqueness half of Aschbacher ex. (8.7)(1) *)
+Lemma isog_pX1p2 (gT : finGroupType) (G : {group gT}) :
+ extraspecial G -> exponent G %| p -> #|G| = (p ^ 3)%N -> G \isog p^{1+2}.
+Proof.
+move=> esG expGp oG; apply/(isoGrpP _ Grp_pX1p2).
+rewrite card_pX1p2; split=> //.
+have pG: p.-group G by rewrite /pgroup oG pnat_exp pnat_id.
+have oZ := card_center_extraspecial pG esG.
+have [x Gx notZx]: exists2 x, x \in G & x \notin 'Z(G).
+ apply/subsetPn; rewrite proper_subn // properEcard center_sub oZ oG.
+ by rewrite (ltn_exp2l 1 3).
+have ox: #[x] = p.
+ by apply: nt_prime_order; rewrite ?(exponentP expGp) ?(group1_contra notZx).
+have [y Gy not_cxy]: exists2 y, y \in G & y \notin 'C[x].
+ by apply/subsetPn; rewrite sub_cent1; rewrite inE Gx in notZx.
+apply/existsP; exists (x, y) => /=; set z := [~ x, y].
+have [[defPhiG defG'] _] := esG.
+have Zz: z \in 'Z(G) by rewrite -defG' mem_commg.
+have [Gz cGz] := setIP Zz; rewrite !xpair_eqE !(exponentP expGp) //.
+have [_ nZG] := andP (center_normal G).
+rewrite /commg /conjg !(centP cGz) // !mulKg mulVg !eqxx !andbT.
+have sXY_G: <[x]> <*> <[y]> \subset G by rewrite join_subG !cycle_subG Gx.
+have defZ: <[z]> = 'Z(G).
+ apply/eqP; rewrite eqEcard cycle_subG Zz oZ /= -orderE.
+ rewrite (nt_prime_order p_pr) ?(exponentP expGp) //.
+ by rewrite (sameP commgP cent1P) cent1C.
+have sZ_XY: 'Z(G) \subset <[x]> <*> <[y]>.
+ by rewrite -defZ cycle_subG groupR // mem_gen // inE cycle_id ?orbT.
+rewrite eqEcard sXY_G /= oG -(Lagrange sZ_XY) oZ leq_pmul2l //.
+rewrite -card_quotient ?(subset_trans sXY_G) //.
+rewrite quotientY ?quotient_cycle ?cycle_subG ?(subsetP nZG) //.
+have abelGz: p.-abelem (G / 'Z(G)) by rewrite -defPhiG Phi_quotient_abelem.
+have [cGzGz expGz] := abelemP p_pr abelGz.
+rewrite cent_joinEr ?(sub_abelian_cent2 cGzGz) ?cycle_subG ?mem_quotient //.
+have oZx: #|<[coset 'Z(G) x]>| = p.
+ rewrite -orderE (nt_prime_order p_pr) ?expGz ?mem_quotient //.
+ by apply: contra notZx; move/eqP=> Zx; rewrite coset_idr ?(subsetP nZG).
+rewrite TI_cardMg ?oZx -?orderE ?(nt_prime_order p_pr) ?expGz ?mem_quotient //.
+ apply: contra not_cxy; move/eqP=> Zy.
+ rewrite -cent_cycle (subsetP _ y (coset_idr _ Zy)) ?(subsetP nZG) //.
+ by rewrite subIset ?centS ?orbT ?cycle_subG.
+rewrite prime_TIg ?oZx // cycle_subG; apply: contra not_cxy.
+case/cycleP=> i; rewrite -morphX ?(subsetP nZG) // => /rcoset_kercosetP.
+rewrite groupX ?(subsetP nZG) // cent1C => /(_ isT isT); apply: subsetP.
+rewrite mul_subG ?sub1set ?groupX ?cent1id //= -cent_cycle subIset // orbC.
+by rewrite centS ?cycle_subG.
+Qed.
+
+End ExponentPextraspecialTheory.
+
+Section GeneralExponentPextraspecialTheory.
+
+Variable p : nat.
+
+Lemma pX1p2id : p^{1+2*1} \isog p^{1+2}.
+Proof. exact: ncprod1. Qed.
+
+Lemma pX1p2S n : xcprod_spec p^{1+2} p^{1+2*n} p^{1+2*n.+1}%type.
+Proof. exact: ncprodS. Qed.
+
+Lemma card_pX1p2n n : prime p -> #|p^{1+2*n}| = (p ^ n.*2.+1)%N.
+Proof.
+move=> p_pr; have pG := pX1p2_pgroup p_pr.
+have oG := card_pX1p2 p_pr; have esG := pX1p2_extraspecial p_pr.
+have oZ := card_center_extraspecial pG esG.
+elim: n => [|n IHn]; first by rewrite (card_isog (ncprod0 _)) oZ.
+case: pX1p2S => gz isoZ; rewrite -im_cpair cardMg_divn setI_im_cpair.
+rewrite -injm_center ?{1}card_injm ?injm_cpairg1 ?injm_cpair1g ?center_sub //.
+by rewrite oG oZ IHn -expnD mulKn ?prime_gt0.
+Qed.
+
+Lemma pX1p2n_pgroup n : prime p -> p.-group p^{1+2*n}.
+Proof. by move=> p_pr; rewrite /pgroup card_pX1p2n // pnat_exp pnat_id. Qed.
+
+(* This is part of the existence half of Aschbacher (23.13) *)
+Lemma exponent_pX1p2n n : prime p -> odd p -> exponent p^{1+2*n} = p.
+Proof.
+move=> p_pr odd_p; apply: prime_nt_dvdP => //.
+ rewrite -dvdn1 -trivg_exponent -cardG_gt1 card_pX1p2n //.
+ by rewrite (ltn_exp2l 0) // prime_gt1.
+elim: n => [|n IHn].
+ by rewrite (dvdn_trans (exponent_dvdn _)) ?card_pX1p2n.
+case: pX1p2S => gz isoZ; rewrite -im_cpair /=.
+apply/exponentP=> xy; case/imset2P=> x y C1x C2y ->{xy}.
+rewrite expgMn; last by red; rewrite -(centsP (im_cpair_cent isoZ)).
+rewrite (exponentP _ y C2y) ?exponent_injm ?injm_cpair1g // mulg1.
+by rewrite (exponentP _ x C1x) ?exponent_injm ?injm_cpairg1 // exponent_pX1p2.
+Qed.
+
+(* This is part of the existence half of Aschbacher (23.13) and (23.14) *)
+Lemma pX1p2n_extraspecial n : prime p -> n > 0 -> extraspecial p^{1+2*n}.
+Proof.
+move=> p_pr; elim: n => [//|n IHn _].
+have esG := pX1p2_extraspecial p_pr.
+have [n0 | n_gt0] := posnP n.
+ by apply: isog_extraspecial esG; rewrite isog_sym n0 pX1p2id.
+case: pX1p2S (pX1p2n_pgroup n.+1 p_pr) => gz isoZ pGn.
+apply: (cprod_extraspecial pGn (im_cpair_cprod isoZ) (setI_im_cpair isoZ)).
+ by apply: injm_extraspecial esG; rewrite ?injm_cpairg1.
+by apply: injm_extraspecial (IHn n_gt0); rewrite ?injm_cpair1g.
+Qed.
+
+(* This is Aschbacher (23.12) *)
+Lemma Ohm1_extraspecial_odd (gT : finGroupType) (G : {group gT}) :
+ p.-group G -> extraspecial G -> odd #|G| ->
+ let Y := 'Ohm_1(G) in
+ [/\ exponent Y = p, #|G : Y| %| p
+ & Y != G ->
+ exists E : {group gT},
+ [/\ #|G : Y| = p, #|E| = p \/ extraspecial E,
+ exists2 X : {group gT}, #|X| = p & X \x E = Y
+ & exists M : {group gT},
+ [/\ M \isog 'Mod_(p ^ 3), M \* E = G & M :&: E = 'Z(M)]]].
+Proof.
+move=> pG esG oddG Y; have [spG _] := esG.
+have [defPhiG defG'] := spG; set Z := 'Z(G) in defPhiG defG'.
+have{spG} expG: exponent G %| p ^ 2 by exact: exponent_special.
+have p_pr := extraspecial_prime pG esG.
+have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1.
+have oZ: #|Z| = p := card_center_extraspecial pG esG.
+have nsZG: Z <| G := center_normal G; have [sZG nZG] := andP nsZG.
+have nsYG: Y <| G := Ohm_normal 1 G; have [sYG nYG] := andP nsYG.
+have ntZ: Z != 1 by rewrite -cardG_gt1 oZ.
+have sZY: Z \subset Y.
+ by apply: contraR ntZ => ?; rewrite -(setIidPl sZG) TI_Ohm1 ?prime_TIg ?oZ.
+have ntY: Y != 1 by apply: contra ntZ; rewrite -!subG1; exact: subset_trans.
+have p_odd: odd p by rewrite -oZ (oddSg sZG).
+have expY: exponent Y %| p by rewrite exponent_Ohm1_class2 // nil_class2 defG'.
+rewrite (prime_nt_dvdP p_pr _ expY) -?dvdn1 -?trivg_exponent //.
+have [-> | neYG] := eqVneq Y G; first by rewrite indexgg dvd1n eqxx; split.
+have sG1Z: 'Mho^1(G) \subset Z by rewrite -defPhiG (Phi_joing pG) joing_subr.
+have Z_Gp: {in G, forall x, x ^+ p \in Z}.
+ by move=> x Gx; rewrite /= (subsetP sG1Z) ?(Mho_p_elt 1) ?(mem_p_elt pG).
+have{expG} oY': {in G :\: Y, forall u, #[u] = (p ^ 2)%N}.
+ move=> u /setDP[Gu notYu]; apply/eqP.
+ have [k ou] := p_natP (mem_p_elt pG Gu).
+ rewrite eqn_dvd order_dvdn (exponentP expG) // eqxx ou dvdn_Pexp2l // ltnNge.
+ apply: contra notYu => k_le_1; rewrite [Y](OhmE _ pG) mem_gen // !inE Gu /=.
+ by rewrite -order_dvdn ou dvdn_exp2l.
+have isoMod3 (M : {group gT}):
+ M \subset G -> ~~ abelian M -> ~~ (M \subset Y) -> #|M| = (p ^ 3)%N ->
+ M \isog 'Mod_(p ^ 3).
+- move=> sMG not_cMM /subsetPn[u Mu notYu oM].
+ have pM := pgroupS sMG pG; have sUM: <[u]> \subset M by rewrite cycle_subG.
+ have Y'u: u \in G :\: Y by rewrite inE notYu (subsetP sMG).
+ have iUM: #|M : <[u]>| = p by rewrite -divgS // oM expnS -(oY' u) ?mulnK.
+ have cM := maximal_cycle_extremal pM not_cMM (cycle_cyclic u) sUM iUM.
+ rewrite (sameP eqP (prime_oddPn p_pr)) p_odd orbF in cM.
+ rewrite /extremal_class oM pdiv_pfactor // pfactorK //= in cM.
+ by do 3!case: ifP => // _ in cM.
+have iYG: #|G : Y| = p.
+ have [V maxV sYV]: {V : {group gT} | maximal V G & Y \subset V}.
+ by apply: maxgroup_exists; rewrite properEneq neYG.
+ have [sVG [u Gu notVu]] := properP (maxgroupp maxV).
+ without loss [v Vv notYv]: / exists2 v, v \in V & v \notin Y.
+ have [->| ] := eqVneq Y V; first by rewrite (p_maximal_index pG).
+ by rewrite eqEsubset sYV => not_sVY; apply; exact/subsetPn.
+ pose U := <[u]> <*> <[v]>; have Gv := subsetP sVG v Vv.
+ have sUG: U \subset G by rewrite join_subG !cycle_subG Gu.
+ have Uu: u \in U by rewrite -cycle_subG joing_subl.
+ have Uv: v \in U by rewrite -cycle_subG joing_subr.
+ have not_sUY: ~~ (U \subset Y) by apply/subsetPn; exists v.
+ have sU1U: 'Ohm_1(U) \subset U := Ohm_sub 1 _.
+ have sU1Y: 'Ohm_1(U) \subset Y := OhmS 1 sUG.
+ suffices defUV: U :&: V = 'Ohm_1(U).
+ by rewrite (subsetP sU1Y) // -defUV inE Uv in notYv.
+ suffices iU1U: #|U : 'Ohm_1(U)| = p.
+ have: maximal 'Ohm_1(U) U by rewrite p_index_maximal ?Ohm_sub ?iU1U.
+ case/maxgroupP=> _; apply; rewrite /= -/U.
+ by apply/properP; split; last exists u; rewrite ?subsetIl ?inE ?Uu.
+ by rewrite subsetI Ohm_sub (subset_trans sU1Y).
+ apply/prime_nt_dvdP=> //.
+ by apply: contra not_sUY; rewrite /U; move/eqP; move/(index1g sU1U)=> <-.
+ have ov: #[v] = (p ^ 2)%N by rewrite oY' // inE notYv.
+ have sZv: Z \subset <[v]>.
+ suffices defZ: <[v ^+ p]> == Z by rewrite -(eqP defZ) cycleX.
+ by rewrite eqEcard cycle_subG Z_Gp //= oZ -orderE (orderXexp 1 ov).
+ have nvG: G \subset 'N(<[v]>) by rewrite sub_der1_norm ?cycle_subG // defG'.
+ have [cUU | not_cUU] := orP (orbN (abelian U)).
+ rewrite -divgS ?Ohm_sub // -(mul_card_Ohm_Mho_abelian 1 cUU) /= -/U.
+ by rewrite mulKn ?cardG_gt0 //= -oZ cardSg ?(subset_trans (MhoS 1 sUG)).
+ have oU: #|U| = (p ^ 3)%N.
+ have nvu := subsetP nvG u Gu; have nvU := subset_trans sUG nvG.
+ rewrite -(Lagrange (joing_subr _ _)) -orderE ov mulnC; congr (_ * _)%N.
+ rewrite -card_quotient //= quotientYidr ?cycle_subG //=.
+ rewrite quotient_cycle // -orderE; apply: nt_prime_order => //.
+ by rewrite -morphX //= coset_id // (subsetP sZv) // Z_Gp.
+ have svV: <[v]> \subset V by rewrite cycle_subG.
+ by apply: contra notVu; move/eqP=> v_u; rewrite (subsetP svV) // coset_idr.
+ have isoU := isoMod3 _ sUG not_cUU not_sUY oU; rewrite /= -/U in isoU.
+ have [//|[x y] genU modU] := generators_modular_group p_pr _ isoU.
+ case/modular_group_structure: genU => // _ _ _ _.
+ case: eqP (p_odd) => [[-> //] | _ _]; case/(_ 1%N)=> // _ oU1.
+ by rewrite -divgS // oU oU1 mulnK // muln_gt0 p_gt0.
+have iC1U (U : {group gT}) x:
+ U \subset G -> x \in G :\: 'C(U) -> #|U : 'C_U[x]| = p.
+- move=> sUG /setDP[Gx not_cUx]; apply/prime_nt_dvdP=> //.
+ apply: contra not_cUx; rewrite -sub_cent1 => /eqP sUCx.
+ by rewrite -(index1g _ sUCx) ?subsetIl ?subsetIr.
+ rewrite -(@dvdn_pmul2l (#|U| * #|'C_G[x]|)) ?muln_gt0 ?cardG_gt0 //.
+ have maxCx: maximal 'C_G[x] G.
+ rewrite cent1_extraspecial_maximal //; apply: contra not_cUx.
+ by rewrite inE Gx; exact: subsetP (centS sUG) _.
+ rewrite {1}mul_cardG setIA (setIidPl sUG) -(p_maximal_index pG maxCx) -!mulnA.
+ rewrite !Lagrange ?subsetIl // mulnC dvdn_pmul2l //.
+ have [sCxG nCxG] := andP (p_maximal_normal pG maxCx).
+ by rewrite -norm_joinEl ?cardSg ?join_subG ?(subset_trans sUG).
+have oCG (U : {group gT}):
+ Z \subset U -> U \subset G -> #|'C_G(U)| = (p * #|G : U|)%N.
+- elim: {U}_.+1 {-2}U (ltnSn #|U|) => // m IHm U leUm sZU sUG.
+ have [<- | neZU] := eqVneq Z U.
+ by rewrite -oZ Lagrange // (setIidPl _) // centsC subsetIr.
+ have{neZU} [x Gx not_cUx]: exists2 x, x \in G & x \notin 'C(U).
+ by apply/subsetPn; rewrite eqEsubset sZU subsetI sUG centsC in neZU.
+ pose W := 'C_U[x]; have iWU: #|U : W| = p by rewrite iC1U // inE not_cUx.
+ have maxW: maximal W U by rewrite p_index_maximal ?subsetIl ?iWU.
+ have ltWU: W \proper U by exact: maxgroupp maxW.
+ have [sWU [u Uu notWu]] := properP ltWU.
+ have defU: W * <[u]> = U.
+ have nsWU: W <| U := p_maximal_normal (pgroupS sUG pG) maxW.
+ by rewrite (mulg_normal_maximal nsWU) ?cycle_subG.
+ have sWG := subset_trans sWU sUG.
+ have sZW: Z \subset W.
+ by rewrite subsetI sZU -cent_set1 subIset ?centS ?orbT ?sub1set.
+ have iCW_CU: #|'C_G(W) : 'C_G(U)| = p.
+ rewrite -defU centM cent_cycle setIA /= -/W.
+ rewrite iC1U ?subsetIl ?setIS ?centS // inE andbC (subsetP sUG) //=.
+ rewrite -sub_cent1; apply/subsetPn; exists x.
+ by rewrite inE Gx -sub_cent1 subsetIr.
+ by rewrite -defU centM cent_cycle inE -sub_cent1 subsetIr in not_cUx.
+ apply/eqP; rewrite -(eqn_pmul2r p_gt0) -{1}iCW_CU Lagrange ?setIS ?centS //.
+ rewrite IHm ?(leq_trans (proper_card ltWU)) //= -/W.
+ by rewrite -(Lagrange_index sUG sWU) iWU mulnA.
+have oCY: #|'C_G(Y)| = (p ^ 2)%N by rewrite oCG // iYG.
+have [x cYx notZx]: exists2 x, x \in 'C_G(Y) & x \notin Z.
+ apply/subsetPn; rewrite proper_subn // properEcard setIS ?centS //=.
+ by rewrite oZ oCY (ltn_exp2l 1 2).
+have{cYx} [Gx cYx] := setIP cYx; have nZx := subsetP nZG x Gx.
+have defCx: 'C_G[x] = Y.
+ apply/eqP; rewrite eq_sym eqEcard subsetI sYG sub_cent1 cYx /=.
+ rewrite -(leq_pmul2r p_gt0) -{2}iYG -(iC1U G x) ?Lagrange ?subsetIl //.
+ by rewrite !inE Gx ?andbT in notZx *.
+have Yx: x \in Y by rewrite -defCx inE Gx cent1id.
+have ox: #[x] = p.
+ by apply: nt_prime_order; rewrite ?(exponentP expY) // (group1_contra notZx).
+have defCy: 'C_G(Y) = Z * <[x]>.
+ apply/eqP; rewrite eq_sym eqEcard mulG_subG setIS ?centS //=.
+ rewrite cycle_subG inE Gx cYx oCY TI_cardMg ?oZ -?orderE ?ox //=.
+ by rewrite setIC prime_TIg -?orderE ?ox ?cycle_subG.
+have abelYt: p.-abelem (Y / Z).
+ by rewrite (abelemS (quotientS _ sYG)) //= -/Z -defPhiG Phi_quotient_abelem.
+have Yxt: coset Z x \in Y / Z by rewrite mem_quotient.
+have{Yxt} [Et [sEtYt oEt defYt]] := p_abelem_split1 abelYt Yxt.
+have nsZY: Z <| Y := normalS sZY sYG nsZG.
+have [E defEt sZE sEY] := inv_quotientS nsZY sEtYt.
+have{defYt} [_ defYt _ tiXEt] := dprodP defYt.
+have defY: <[x]> \x E = Y.
+ have nZX: <[x]> \subset 'N(Z) by rewrite cycle_subG.
+ have TIxE: <[x]> :&: E = 1.
+ rewrite prime_TIg -?orderE ?ox // -(quotientSGK _ sZE) ?quotient_cycle //.
+ rewrite (sameP setIidPl eqP) eq_sym -defEt tiXEt -quotient_cycle //.
+ by rewrite -subG1 quotient_sub1 // cycle_subG.
+ rewrite dprodE //; last 1 first.
+ by rewrite cent_cycle (subset_trans sEY) //= -/Y -defCx subsetIr.
+ rewrite -[Y](quotientGK nsZY) -defYt cosetpreM -quotient_cycle //.
+ rewrite quotientK // -(normC nZX) defEt quotientGK ?(normalS _ sEY) //.
+ by rewrite -mulgA (mulSGid sZE).
+have sEG := subset_trans sEY sYG; have nZE := subset_trans sEG nZG.
+have defZE: 'Z(E) = Z.
+ apply/eqP; rewrite eqEsubset andbC subsetI sZE subIset ?centS ?orbT //.
+ rewrite -quotient_sub1 ?subIset ?nZE //= -tiXEt defEt subsetI andbC.
+ rewrite quotientS ?center_sub //= -quotient_cycle //.
+ rewrite -(quotientMidl _ <[x]>) /= -defCy quotientS // /Y.
+ by case/dprodP: defY => _ <- _ _; rewrite centM setIA cent_cycle defCx setSI.
+have pE := pgroupS sEG pG.
+rewrite iYG; split=> // _; exists E.
+split=> //; first 2 [by exists [group of <[x]>]].
+ have:= sZE; rewrite subEproper; case/predU1P=> [<- | ltZE]; [by left | right].
+ split; rewrite /special defZE ?oZ // (Phi_joing pE).
+ have defE': E^`(1) = Z.
+ have sE'Z: E^`(1) \subset Z by rewrite -defG' dergS.
+ apply/eqP; rewrite eqEcard sE'Z -(prime_nt_dvdP _ _ (cardSg sE'Z)) ?oZ //=.
+ rewrite -trivg_card1 (sameP eqP commG1P).
+ by rewrite /proper sZE /= -/Z -defZE subsetI subxx in ltZE.
+ split=> //; rewrite -defE'; apply/joing_idPl.
+ by rewrite /= defE' -defPhiG (Phi_joing pG) joingC sub_gen ?subsetU ?MhoS.
+have iEG: #|G : E| = (p ^ 2)%N.
+ apply/eqP; rewrite -(@eqn_pmul2l #|E|) // Lagrange // -(Lagrange sYG) iYG.
+ by rewrite -(dprod_card defY) -mulnA mulnCA -orderE ox.
+pose M := 'C_G(E); exists [group of M] => /=.
+have sMG: M \subset G := subsetIl _ _; have pM: p.-group M := pgroupS sMG pG.
+have sZM: Z \subset M by rewrite setIS ?centS.
+have oM: #|M| = (p ^ 3)%N by rewrite oCG ?iEG.
+have defME: M * E = G.
+ apply/eqP; rewrite eqEcard mulG_subG sMG sEG /= -(leq_pmul2r p_gt0).
+ rewrite -{2}oZ -defZE /('Z(E)) -{2}(setIidPr sEG) setIAC -mul_cardG /= -/M.
+ by rewrite -(Lagrange sEG) mulnAC -mulnA mulnC iEG oM.
+have defZM: 'Z(M) = Z.
+ apply/eqP; rewrite eqEsubset andbC subsetI sZM subIset ?centS ?orbT //=.
+ by rewrite /Z /('Z(G)) -{2}defME centM setIA setIAC.
+rewrite cprodE 1?centsC ?subsetIr //.
+rewrite defME setIAC (setIidPr sEG) defZM isoMod3 //.
+ rewrite abelianE (sameP setIidPl eqP) eqEcard subsetIl /= -/('Z(M)) -/M.
+ by rewrite defZM oZ oM (leq_exp2l 3 1).
+by apply: contra neYG => sMY; rewrite eqEsubset sYG -defME mulG_subG sMY.
+Qed.
+
+(* This is the uniqueness half of Aschbacher (23.13); the proof incorporates *)
+(* in part the proof that symplectic spaces are hyperbolic (19.16). *)
+Lemma isog_pX1p2n n (gT : finGroupType) (G : {group gT}) :
+ prime p -> extraspecial G -> #|G| = (p ^ n.*2.+1)%N -> exponent G %| p ->
+ G \isog p^{1+2*n}.
+Proof.
+move=> p_pr esG oG expG; have p_gt1 := prime_gt1 p_pr.
+have not_le_p3_p: ~~ (p ^ 3 <= p) by rewrite (leq_exp2l 3 1).
+have pG: p.-group G by rewrite /pgroup oG pnat_exp pnat_id.
+have oZ := card_center_extraspecial pG esG.
+have{pG esG} [Es p3Es defG] := extraspecial_structure pG esG.
+set Z := 'Z(G) in oZ defG p3Es.
+elim: Es {+}G => [|E Es IHs] S in n oG expG p3Es defG *.
+ rewrite big_nil cprod1g in defG; rewrite -defG.
+ have ->: n = 0%N.
+ apply: double_inj; apply/eqP.
+ by rewrite -eqSS -(eqn_exp2l _ _ p_gt1) -oG -defG oZ.
+ by rewrite isog_cyclic_card prime_cyclic ?oZ ?card_pX1p2n //=.
+rewrite big_cons -cprodA in defG; rewrite /= -andbA in p3Es.
+have [[_ T _ defT] defET cTE] := cprodP defG; rewrite defT in defET cTE defG.
+case/and3P: p3Es; move/eqP=> oE; move/eqP=> defZE; move/IHs=> {IHs}IHs.
+have not_cEE: ~~ abelian E.
+ by apply: contra not_le_p3_p => cEE; rewrite -oE -oZ -defZE (center_idP _).
+have sES: E \subset S by rewrite -defET mulG_subl.
+have sTS: T \subset S by rewrite -defET mulG_subr.
+have expE: exponent E %| p by exact: dvdn_trans (exponentS sES) expG.
+have expT: exponent T %| p by exact: dvdn_trans (exponentS sTS) expG.
+have{expE not_cEE} isoE: E \isog p^{1+2}.
+ apply: isog_pX1p2 => //.
+ by apply: card_p3group_extraspecial p_pr oE _; rewrite defZE.
+have sZT: 'Z(E) \subset T.
+ by case/cprodP: defT => [[U _ -> _] <- _]; rewrite defZE mulG_subr.
+case def_n: n => [|n'].
+ case/negP: not_le_p3_p; rewrite def_n in oG; rewrite -oE -[p]oG.
+ exact: subset_leq_card.
+have zI_ET: E :&: T = 'Z(E).
+ by apply/eqP; rewrite eqEsubset subsetI sZT subsetIl setIS // centsC.
+have{n def_n oG} oT: #|T| = (p ^ n'.*2.+1)%N.
+ apply/eqP; rewrite -(eqn_pmul2l (cardG_gt0 E)) mul_cardG zI_ET defET.
+ by rewrite defZE oE oG oZ -expnSr -expnD def_n.
+have{IHs oT expT defT Es} isoT: T \isog p^{1+2*n'} by rewrite IHs.
+case: pX1p2S => gz isoZ; rewrite (isog_cprod_by _ defG) //.
+exact: Aut_extraspecial_full (pX1p2_pgroup p_pr) (pX1p2_extraspecial p_pr).
+Qed.
+
+End GeneralExponentPextraspecialTheory.
+
+Lemma isog_2X1p2 : 2^{1+2} \isog 'D_8.
+Proof.
+have pr2: prime 2 by []; have oG := card_pX1p2 pr2; rewrite -[8]oG.
+case/existsP: (isoGrp_hom (Grp_pX1p2 pr2)) => [[x y]] /=.
+rewrite -/2^{1+2}; case/eqP=> defG x2 y2 _ _.
+have not_oG_2: ~~ (#|2^{1+2}| %| 2) by rewrite oG.
+have ox: #[x] = 2.
+ apply: nt_prime_order => //; apply: contra not_oG_2 => x1.
+ by rewrite -defG (eqP x1) cycle1 joing1G order_dvdn y2.
+have oy: #[y] = 2.
+ apply: nt_prime_order => //; apply: contra not_oG_2 => y1.
+ by rewrite -defG (eqP y1) cycle1 joingG1 order_dvdn x2.
+rewrite -defG joing_idl joing_idr involutions_gen_dihedral //.
+apply: contra not_oG_2 => eq_xy; rewrite -defG (eqP eq_xy) (joing_idPl _) //.
+by rewrite -orderE oy.
+Qed.
+
+Lemma Q8_extraspecial : extraspecial 'Q_8.
+Proof.
+have gt32: 3 > 2 by []; have isoQ: 'Q_8 \isog 'Q_(2 ^ 3) by exact: isog_refl.
+have [[x y] genQ _] := generators_quaternion gt32 isoQ.
+have [_ [defQ' defPhiQ _ _]] := quaternion_structure gt32 genQ isoQ.
+case=> defZ oZ _ _ _ _ _; split; last by rewrite oZ.
+by split; rewrite ?defPhiQ defZ.
+Qed.
+
+Lemma DnQ_P n : xcprod_spec 'D^n 'Q_8 ('D^n*Q)%type.
+Proof.
+have pQ: 2.-group 'Q_(2 ^ 3) by rewrite /pgroup card_quaternion.
+have{pQ} oZQ := card_center_extraspecial pQ Q8_extraspecial.
+suffices oZDn: #|'Z('D^n)| = 2.
+ by apply: xcprodP; rewrite isog_cyclic_card ?prime_cyclic ?oZQ ?oZDn.
+have [-> | n_gt0] := posnP n; first by rewrite center_ncprod0 card_pX1p2n.
+have pr2: prime 2 by []; have pDn := pX1p2n_pgroup n pr2.
+exact: card_center_extraspecial (pX1p2n_extraspecial pr2 n_gt0).
+Qed.
+
+Lemma card_DnQ n : #|'D^n*Q| = (2 ^ n.+1.*2.+1)%N.
+Proof.
+have oQ: #|'Q_(2 ^ 3)| = 8 by rewrite card_quaternion.
+have pQ: 2.-group 'Q_8 by rewrite /pgroup oQ.
+case: DnQ_P => gz isoZ.
+rewrite -im_cpair cardMg_divn setI_im_cpair cpair_center_id.
+rewrite -injm_center 3?{1}card_injm ?injm_cpairg1 ?injm_cpair1g ?center_sub //.
+rewrite oQ card_pX1p2n // (card_center_extraspecial pQ Q8_extraspecial).
+by rewrite -muln_divA // mulnC -(expnD 2 2).
+Qed.
+
+Lemma DnQ_pgroup n : 2.-group 'D^n*Q.
+Proof. by rewrite /pgroup card_DnQ pnat_exp. Qed.
+
+(* Final part of the existence half of Aschbacher (23.14). *)
+Lemma DnQ_extraspecial n : extraspecial 'D^n*Q.
+Proof.
+case: DnQ_P (DnQ_pgroup n) => gz isoZ pDnQ.
+have [injDn injQ] := (injm_cpairg1 isoZ, injm_cpair1g isoZ).
+have [n0 | n_gt0] := posnP n.
+ rewrite -im_cpair mulSGid; first exact: injm_extraspecial Q8_extraspecial.
+ apply/setIidPl; rewrite setI_im_cpair -injm_center //=.
+ by congr (_ @* _); rewrite n0 center_ncprod0.
+apply: (cprod_extraspecial pDnQ (im_cpair_cprod isoZ) (setI_im_cpair _)).
+ exact: injm_extraspecial (pX1p2n_extraspecial _ _).
+exact: injm_extraspecial Q8_extraspecial.
+Qed.
+
+(* A special case of the uniqueness half of Achsbacher (23.14). *)
+Lemma card_isog8_extraspecial (gT : finGroupType) (G : {group gT}) :
+ #|G| = 8 -> extraspecial G -> (G \isog 'D_8) || (G \isog 'Q_8).
+Proof.
+move=> oG esG; have pG: 2.-group G by rewrite /pgroup oG.
+apply/norP=> [[notG_D8 notG_Q8]].
+have not_extG: extremal_class G = NotExtremal.
+ by rewrite /extremal_class oG andFb (negPf notG_D8) (negPf notG_Q8).
+have [x Gx ox] := exponent_witness (pgroup_nil pG).
+pose X := <[x]>; have cycX: cyclic X := cycle_cyclic x.
+have sXG: X \subset G by rewrite cycle_subG.
+have iXG: #|G : X| = 2.
+ by rewrite -divgS // oG -orderE -ox exponent_2extraspecial.
+have not_cGG := extraspecial_nonabelian esG.
+have:= maximal_cycle_extremal pG not_cGG cycX sXG iXG.
+by rewrite /extremal2 not_extG.
+Qed.
+
+(* The uniqueness half of Achsbacher (23.14). The proof incorporates in part *)
+(* the proof that symplectic spces are hyperbolic (Aschbacher (19.16)), and *)
+(* the determination of quadratic spaces over 'F_2 (21.2); however we use *)
+(* the second part of exercise (8.4) to avoid resorting to Witt's lemma and *)
+(* Galois theory as in (20.9) and (21.1). *)
+Lemma isog_2extraspecial (gT : finGroupType) (G : {group gT}) n :
+ #|G| = (2 ^ n.*2.+1)%N -> extraspecial G -> G \isog 'D^n \/ G \isog 'D^n.-1*Q.
+Proof.
+elim: n G => [|n IHn] G oG esG.
+ case/negP: (extraspecial_nonabelian esG).
+ by rewrite cyclic_abelian ?prime_cyclic ?oG.
+have pG: 2.-group G by rewrite /pgroup oG pnat_exp.
+have oZ:= card_center_extraspecial pG esG.
+have: 'Z(G) \subset 'Ohm_1(G).
+ apply/subsetP=> z Zz; rewrite (OhmE _ pG) mem_gen //.
+ by rewrite !inE -order_dvdn -oZ order_dvdG ?(subsetP (center_sub G)).
+rewrite subEproper; case/predU1P=> [defG1 | ltZG1].
+ have [n' n'_gt2 isoG]: exists2 n', n' > 2 & G \isog 'Q_(2 ^ n').
+ apply/quaternion_classP; apply/eqP.
+ have not_cycG: ~~ cyclic G.
+ by apply: contra (extraspecial_nonabelian esG); exact: cyclic_abelian.
+ move: oZ; rewrite defG1; move/prime_Ohm1P; rewrite (negPf not_cycG) /=.
+ by apply=> //; apply: contra not_cycG; move/eqP->; exact: cyclic1.
+ have [n0 n'3]: n = 0%N /\ n' = 3.
+ have [[x y] genG _] := generators_quaternion n'_gt2 isoG.
+ have n'3: n' = 3.
+ have [_ [_ _ oG' _] _ _ _] := quaternion_structure n'_gt2 genG isoG.
+ apply/eqP; rewrite -(subnKC (ltnW n'_gt2)) subn2 !eqSS -(@eqn_exp2l 2) //.
+ by rewrite -oG' -oZ; case: esG => [[_ ->]].
+ by move/eqP: oG; have [-> _ _ _] := genG; rewrite n'3 eqn_exp2l //; case n.
+ right; rewrite (isog_trans isoG) // n'3 n0 /=.
+ case: DnQ_P => z isoZ; rewrite -im_cpair mulSGid ?sub_isog ?injm_cpair1g //.
+ apply/setIidPl; rewrite setI_im_cpair -injm_center ?injm_cpairg1 //.
+ by rewrite center_ncprod0.
+case/andP: ltZG1 => _; rewrite (OhmE _ pG) gen_subG.
+case/subsetPn=> x; case/LdivP=> Gx x2 notZx.
+have ox: #[x] = 2 by exact: nt_prime_order (group1_contra notZx).
+have Z'x: x \in G :\: 'Z(G) by rewrite inE notZx.
+have [E [R [[oE oR] [defG ziER]]]] := split1_extraspecial pG esG Z'x.
+case=> defZE defZR [esE Ex] esR.
+have isoE: E \isog 2^{1+2}.
+ apply: isog_trans (isog_symr isog_2X1p2).
+ case/orP: (card_isog8_extraspecial oE esE) => // isoE; case/negP: notZx.
+ have gt32: 3 > 2 by [].
+ have [[y z] genE _] := generators_quaternion gt32 isoE.
+ have [_ _ [defZx _ eq_y2 _ _] _ _] := quaternion_structure gt32 genE isoE.
+ by rewrite (eq_y2 x) // -cycle_subG -defZx defZE.
+rewrite oG doubleS 2!expnS divnMl ?mulKn // in oR.
+case: ifP esR => [_ defR | _ esR].
+ have ->: n = 0%N by move/eqP: oR; rewrite defR oZ (eqn_exp2l 1) //; case n.
+ left; apply: isog_trans (isog_symr (ncprod1 _)).
+ by rewrite -defG defR -defZE cprod_center_id.
+have AutZin2_1p2: Aut_in (Aut 2^{1+2}) 'Z(2^{1+2}) \isog Aut 'Z(2^{1+2}).
+ exact: Aut_extraspecial_full (pX1p2_pgroup _) (pX1p2_extraspecial _).
+have [isoR | isoR] := IHn R oR esR.
+ by left; case: pX1p2S => gz isoZ; rewrite (isog_cprod_by _ defG).
+have n_gt0: n > 0.
+ have pR: 2.-group R by rewrite /pgroup oR pnat_exp.
+ have:= min_card_extraspecial pR esR.
+ by rewrite oR leq_exp2l // ltnS (leq_double 1).
+case: DnQ_P isoR => gR isoZR /=; rewrite isog_sym; case/isogP=> fR injfR im_fR.
+have [injDn injQ] := (injm_cpairg1 isoZR, injm_cpair1g isoZR).
+pose Dn1 := cpairg1 isoZR @* 'D^n.-1; pose Q := cpair1g isoZR @* 'Q_8.
+have defR: fR @* Dn1 \* fR @* Q = R.
+ rewrite cprodE ?morphim_cents ?im_cpair_cent //.
+ by rewrite -morphimMl ?subsetT ?im_cpair.
+rewrite -defR cprodA in defG.
+have [[Dn _ defDn _] _ _] := cprodP defG; rewrite defDn in defG.
+have isoDn: Dn \isog 'D^n.
+ rewrite -(prednK n_gt0); case: pX1p2S => gz isoZ.
+ rewrite (isog_cprod_by _ defDn) //; last 1 first.
+ by rewrite isog_sym (isog_trans _ (sub_isog _ _)) ?subsetT // sub_isog.
+ rewrite /= -morphimIim im_fR setIA ziER; apply/setIidPl.
+ rewrite defZE -defZR -{1}im_fR -injm_center // morphimS //.
+ by rewrite -cpairg1_center morphimS // center_sub.
+right; case: DnQ_P => gz isoZ; rewrite (isog_cprod_by _ defG) //; first 1 last.
+- exact: Aut_extraspecial_full (pX1p2n_pgroup _ _) (pX1p2n_extraspecial _ _).
+- by rewrite isog_sym (isog_trans _ (sub_isog _ _)) ?subsetT // sub_isog.
+rewrite /= -morphimIim; case/cprodP: defDn => _ defDn cDn1E.
+rewrite setICA setIA -defDn -group_modr ?morphimS ?subsetT //.
+rewrite /= im_fR (setIC R) ziER -center_prod // defZE -defZR.
+rewrite mulSGid /=; last first.
+ by rewrite -{1}im_fR -injm_center // -cpairg1_center !morphimS ?center_sub.
+rewrite -injm_center ?subsetT // -injmI // setI_im_cpair.
+by rewrite -injm_center // cpairg1_center injm_center // im_fR mulGid.
+Qed.
+
+(* The first concluding remark of Aschbacher (23.14). *)
+Lemma rank_Dn n : 'r_2('D^n) = n.+1.
+Proof.
+elim: n => [|n IHn]; first by rewrite p_rank_abelem ?prime_abelem ?card_pX1p2n.
+have oDDn: #|'D^n.+1| = (2 ^ n.+1.*2.+1)%N by exact: card_pX1p2n.
+have esDDn: extraspecial 'D^n.+1 by exact: pX1p2n_extraspecial.
+do [case: pX1p2S => gz isoZ; set DDn := [set: _]] in oDDn esDDn *.
+have pDDn: 2.-group DDn by rewrite /pgroup oDDn pnat_exp.
+apply/eqP; rewrite eqn_leq; apply/andP; split.
+ have [E EprE]:= p_rank_witness 2 [group of DDn].
+ have [sEDDn abelE <-] := pnElemP EprE; have [pE cEE _]:= and3P abelE.
+ rewrite -(@leq_exp2l 2) // -p_part part_pnat_id // -leq_sqr -expnM -mulnn.
+ rewrite muln2 doubleS expnS -oDDn -(@leq_pmul2r #|'C_DDn(E)|) ?cardG_gt0 //.
+ rewrite {1}(card_subcent_extraspecial pDDn) // mulnCA -mulnA Lagrange //=.
+ rewrite mulnAC mulnA leq_pmul2r ?cardG_gt0 // setTI.
+ have ->: (2 * #|'C(E)| = #|'Z(DDn)| * #|'C(E)|)%N.
+ by rewrite (card_center_extraspecial pDDn).
+ by rewrite leq_mul ?subset_leq_card ?subsetIl.
+have [inj1 injn] := (injm_cpairg1 isoZ, injm_cpair1g isoZ).
+pose D := cpairg1 isoZ @* 2^{1+2}; pose Dn := cpair1g isoZ @* 'D^n.
+have [E EprE] := p_rank_witness 2 [group of Dn].
+rewrite injm_p_rank //= IHn in EprE; have [sEDn abelE dimE]:= pnElemP EprE.
+have [x [Dx ox] notDnx]: exists x, [/\ x \in D, #[x] = 2 & x \notin Dn].
+ have isoD: D \isog 'D_(2 ^ 3).
+ by rewrite isog_sym -(isog_transl _ isog_2X1p2) sub_isog.
+ have [//| [x y] genD [oy _]] := generators_2dihedral _ isoD.
+ have [_ _ _ X'y] := genD; case/setDP: X'y; rewrite /= -/D => Dy notXy.
+ exists y; split=> //; apply: contra notXy => Dny.
+ case/dihedral2_structure: genD => // _ _ _ _ [defZD _ _ _ _].
+ by rewrite (subsetP (cycleX x 2)) // -defZD -setI_im_cpair inE Dy.
+have def_xE: <[x]> \x E = <[x]> <*> E.
+ rewrite dprodEY ?prime_TIg -?orderE ?ox //.
+ by rewrite (centSS sEDn _ (im_cpair_cent _)) ?cycle_subG.
+ by rewrite cycle_subG (contra (subsetP sEDn x)).
+apply/p_rank_geP; exists (<[x]> <*> E)%G.
+rewrite 2!inE subsetT (dprod_abelem _ def_xE) abelE -(dprod_card def_xE).
+by rewrite prime_abelem -?orderE ?ox //= lognM ?cardG_gt0 ?dimE.
+Qed.
+
+(* The second concluding remark of Aschbacher (23.14). *)
+Lemma rank_DnQ n : 'r_2('D^n*Q) = n.+1.
+Proof.
+have pDnQ: 2.-group 'D^n*Q := DnQ_pgroup n.
+have esDnQ: extraspecial 'D^n*Q := DnQ_extraspecial n.
+do [case: DnQ_P => gz isoZ; set DnQ := setT] in pDnQ esDnQ *.
+suffices [E]: exists2 E, E \in 'E*_2(DnQ) & logn 2 #|E| = n.+1.
+ by rewrite (pmaxElem_extraspecial pDnQ esDnQ); case/pnElemP=> _ _ <-.
+have oZ: #|'Z(DnQ)| = 2 by exact: card_center_extraspecial.
+pose Dn := cpairg1 isoZ @* 'D^n; pose Q := cpair1g isoZ @* 'Q_8.
+have [injDn injQ] := (injm_cpairg1 isoZ, injm_cpair1g isoZ).
+have [E EprE]:= p_rank_witness 2 [group of Dn].
+have [sEDn abelE dimE] := pnElemP EprE; have [pE cEE eE]:= and3P abelE.
+rewrite injm_p_rank // rank_Dn in dimE; exists E => //.
+have sZE: 'Z(DnQ) \subset E.
+ have maxE := subsetP (p_rankElem_max _ _) E EprE.
+ have abelZ: 2.-abelem 'Z(DnQ) by rewrite prime_abelem ?oZ.
+ rewrite -(Ohm1_id abelZ) (OhmE _ (abelem_pgroup abelZ)) gen_subG.
+ rewrite -(pmaxElem_LdivP _ maxE) // setSI //=.
+ by rewrite -cpairg1_center injm_center // setIS ?centS.
+have scE: 'C_Dn(E) = E.
+ apply/eqP; rewrite eq_sym eqEcard subsetI sEDn -abelianE cEE /=.
+ have [n0 | n_gt0] := posnP n.
+ rewrite subset_leq_card // subIset // (subset_trans _ sZE) //.
+ by rewrite -cpairg1_center morphimS // n0 center_ncprod0.
+ have pDn: 2.-group Dn by rewrite morphim_pgroup ?pX1p2n_pgroup.
+ have esDn: extraspecial Dn.
+ exact: injm_extraspecial (pX1p2n_extraspecial _ _).
+ rewrite dvdn_leq ?cardG_gt0 // (card_subcent_extraspecial pDn) //=.
+ rewrite -injm_center // cpairg1_center (setIidPl sZE) oZ.
+ rewrite -(dvdn_pmul2l (cardG_gt0 E)) mulnn mulnCA Lagrange //.
+ rewrite card_injm ?card_pX1p2n // -expnS pfactor_dvdn ?expn_gt0 ?cardG_gt0 //.
+ by rewrite lognX dimE mul2n.
+apply/pmaxElemP; split=> [|F E2F sEF]; first by rewrite inE subsetT abelE.
+have{E2F} [_ abelF] := pElemP E2F; have [pF cFF eF] := and3P abelF.
+apply/eqP; rewrite eqEsubset sEF andbT; apply/subsetP=> x Fx.
+have DnQx: x \in Dn * Q by rewrite im_cpair inE.
+have{DnQx} [y z Dn_y Qz def_x]:= imset2P DnQx.
+have{Dn_y} Ey: y \in E.
+ have cEz: z \in 'C(E).
+ by rewrite (subsetP (centS sEDn)) // (subsetP (im_cpair_cent _)).
+ rewrite -scE inE Dn_y -(groupMr _ cEz) -def_x (subsetP (centS sEF)) //.
+ by rewrite (subsetP cFF).
+rewrite def_x groupMl // (subsetP sZE) // -cpair1g_center injm_center //= -/Q.
+have: z \in 'Ohm_1(Q).
+ rewrite (OhmE 1 (pgroupS (subsetT Q) pDnQ)) mem_gen // !inE Qz /=.
+ rewrite -[z](mulKg y) -def_x (exponentP eF) ?groupM //.
+ by rewrite groupV (subsetP sEF).
+have isoQ: Q \isog 'Q_(2 ^ 3) by rewrite isog_sym sub_isog.
+have [//|[u v] genQ _] := generators_quaternion _ isoQ.
+by case/quaternion_structure: genQ => // _ _ [-> _ _ [-> _] _] _ _.
+Qed.
+
+(* The final concluding remark of Aschbacher (23.14). *)
+Lemma not_isog_Dn_DnQ n : ~~ ('D^n \isog 'D^n.-1*Q).
+Proof.
+case: n => [|n] /=; first by rewrite isogEcard card_pX1p2n // card_DnQ andbF.
+apply: contraL (leqnn n.+1) => isoDn1DnQ.
+by rewrite -ltnNge -rank_Dn (isog_p_rank isoDn1DnQ) rank_DnQ.
+Qed.
diff --git a/mathcomp/solvable/extremal.v b/mathcomp/solvable/extremal.v
new file mode 100644
index 0000000..b8cb69d
--- /dev/null
+++ b/mathcomp/solvable/extremal.v
@@ -0,0 +1,2331 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div choice fintype.
+Require Import bigop finset prime binomial fingroup morphism perm automorphism.
+Require Import presentation quotient action commutator gproduct gfunctor.
+Require Import ssralg finalg zmodp cyclic pgroup center gseries.
+Require Import nilpotent sylow abelian finmodule matrix maximal.
+
+(******************************************************************************)
+(* This file contains the definition and properties of extremal p-groups; *)
+(* it covers and is mostly based on the beginning of Aschbacher, section 23, *)
+(* as well as several exercises of this section. *)
+(* We define canonical representatives for the group classes that cover the *)
+(* extremal p-groups (non-abelian p-groups with a cyclic maximal subgroup): *)
+(* 'Mod_m == the modular group of order m, for m = p ^ n, p prime and n >= 3. *)
+(* 'D_m == the dihedral group of order m, for m = 2n >= 4. *)
+(* 'Q_m == the generalized quaternion group of order m, for m = 2 ^ n >= 8. *)
+(* 'SD_m == the semi-dihedral group of order m, for m = 2 ^ n >= 16. *)
+(* In each case the notation is defined in the %type, %g and %G scopes, where *)
+(* it denotes a finGroupType, a full gset and the full group for that type. *)
+(* However each notation is only meaningful under the given conditions, in *)
+(* 'D_m is only an extremal group for m = 2 ^ n >= 8, and 'D_8 = 'Mod_8 (they *)
+(* are, in fact, beta-convertible). *)
+(* We also define *)
+(* extremal_generators G p n (x, y) <-> G has order p ^ n, x in G has order *)
+(* p ^ n.-1, and y is in G \ <[x]>: thus <[x]> has index p in G, *)
+(* so if p is prime, <[x]> is maximal in G, G is generated by x *)
+(* and y, and G is extremal or abelian. *)
+(* extremal_class G == the class of extremal groups G belongs to: one of *)
+(* ModularGroup, Dihedral, Quaternion, SemiDihedral or NotExtremal. *)
+(* extremal2 G <=> extremal_class G is one of Dihedral, Quaternion, or *)
+(* SemiDihedral; this allows 'D_4 and 'D_8, but excludes 'Mod_(2^n) *)
+(* for n > 3. *)
+(* modular_group_generators p n (x, y) <-> y has order p and acts on x via *)
+(* x ^ y = x ^+ (p ^ n.-2).+1. This is the complement to *)
+(* extremal_generators G p n (x, y) for modular groups. *)
+(* We provide cardinality, presentation, generator and structure theorems for *)
+(* each class of extremal group. The extremal_generators predicate is used to *)
+(* supply structure theorems with all the required data about G; this is *)
+(* completed by an isomorphism assumption (e.g., G \isog 'D_(2 ^ n)), and *)
+(* sometimes other properties (e.g., #[y] == 2 in the semidihedral case). The *)
+(* generators assumption can be deduced generically from the isomorphism *)
+(* assumption, or it can be proved manually for a specific choice of x and y. *)
+(* The extremal_class function is used to formulate synthetic theorems that *)
+(* cover several classes of extremal groups (e.g., Aschbacher ex. 8.3). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GroupScope GRing.Theory.
+
+Reserved Notation "''Mod_' m" (at level 8, m at level 2, format "''Mod_' m").
+Reserved Notation "''D_' m" (at level 8, m at level 2, format "''D_' m").
+Reserved Notation "''SD_' m" (at level 8, m at level 2, format "''SD_' m").
+Reserved Notation "''Q_' m" (at level 8, m at level 2, format "''Q_' m").
+
+Module Extremal.
+
+Section Construction.
+
+Variables q p e : nat.
+(* Construct the semi-direct product of 'Z_q by 'Z_p with 1%R ^ 1%R = e%R, *)
+(* if possible, i.e., if p, q > 1 and there is s \in Aut 'Z_p such that *)
+(* #[s] %| p and s 1%R = 1%R ^+ e. *)
+
+Let a : 'Z_p := Zp1.
+Let b : 'Z_q := Zp1.
+Local Notation B := <[b]>.
+
+Definition aut_of :=
+ odflt 1 [pick s in Aut B | p > 1 & (#[s] %| p) && (s b == b ^+ e)].
+
+Lemma aut_dvdn : #[aut_of] %| #[a].
+Proof.
+rewrite order_Zp1 /aut_of; case: pickP => [s | _]; last by rewrite order1.
+by case/and4P=> _ p_gt1 p_s _; rewrite Zp_cast.
+Qed.
+
+Definition act_morphism := eltm_morphism aut_dvdn.
+
+Definition base_act := ([Aut B] \o act_morphism)%gact.
+
+Lemma act_dom : <[a]> \subset act_dom base_act.
+Proof.
+rewrite cycle_subG 2!inE cycle_id /= eltm_id /aut_of.
+by case: pickP => [op /andP[] | _] //=; rewrite group1.
+Qed.
+
+Definition gact := (base_act \ act_dom)%gact.
+Fact gtype_key : unit. Proof. by []. Qed.
+Definition gtype := locked_with gtype_key (sdprod_groupType gact).
+
+Hypotheses (p_gt1 : p > 1) (q_gt1 : q > 1).
+
+Lemma card : #|[set: gtype]| = (p * q)%N.
+Proof.
+rewrite [gtype]unlock -(sdprod_card (sdprod_sdpair _)).
+rewrite !card_injm ?injm_sdpair1 ?injm_sdpair2 //.
+by rewrite mulnC -!orderE !order_Zp1 !Zp_cast.
+Qed.
+
+Lemma Grp : (exists s, [/\ s \in Aut B, #[s] %| p & s b = b ^+ e]) ->
+ [set: gtype] \isog Grp (x : y : (x ^+ q, y ^+ p, x ^ y = x ^+ e)).
+Proof.
+rewrite [gtype]unlock => [[s [AutBs dvd_s_p sb]]].
+have memB: _ \in B by move=> c; rewrite -Zp_cycle inE.
+have Aa: a \in <[a]> by rewrite !cycle_id.
+have [oa ob]: #[a] = p /\ #[b] = q by rewrite !order_Zp1 !Zp_cast.
+have def_s: aut_of = s.
+ rewrite /aut_of; case: pickP => /= [t | ]; last first.
+ by move/(_ s); case/and4P; rewrite sb.
+ case/and4P=> AutBt _ _ tb; apply: (eq_Aut AutBt) => // b_i.
+ case/cycleP=> i ->; rewrite -(autmE AutBt) -(autmE AutBs) !morphX //=.
+ by rewrite !autmE // sb (eqP tb).
+apply: intro_isoGrp => [|gT G].
+ apply/existsP; exists (sdpair1 _ b, sdpair2 _ a); rewrite /= !xpair_eqE.
+ rewrite -!morphim_cycle ?norm_joinEr ?im_sdpair ?im_sdpair_norm ?eqxx //=.
+ rewrite -!order_dvdn !order_injm ?injm_sdpair1 ?injm_sdpair2 // oa ob !dvdnn.
+ by rewrite -sdpair_act // [act _ _ _]apermE /= eltm_id -morphX // -sb -def_s.
+case/existsP=> -[x y] /= /eqP[defG xq1 yp1 xy].
+have fxP: #[x] %| #[b] by rewrite order_dvdn ob xq1.
+have fyP: #[y] %| #[a] by rewrite order_dvdn oa yp1.
+have fP: {in <[b]> & <[a]>, morph_act gact 'J (eltm fxP) (eltm fyP)}.
+ move=> bj ai; case/cycleP=> j ->{bj}; case/cycleP=> i ->{ai}.
+ rewrite /= !eltmE def_s gactX ?groupX // conjXg morphX //=; congr (_ ^+ j).
+ rewrite /autact /= apermE; elim: i {j} => /= [|i IHi].
+ by rewrite perm1 eltm_id conjg1.
+ rewrite !expgS permM sb -(autmE (groupX i AutBs)) !morphX //= {}IHi.
+ by rewrite -conjXg -xy -conjgM.
+apply/homgP; exists (xsdprod_morphism fP).
+rewrite im_xsdprodm !morphim_cycle //= !eltm_id -norm_joinEr //.
+by rewrite norms_cycle xy mem_cycle.
+Qed.
+
+End Construction.
+
+End Extremal.
+
+Section SpecializeExtremals.
+
+Import Extremal.
+
+Variable m : nat.
+Let p := pdiv m.
+Let q := m %/ p.
+
+Definition modular_gtype := gtype q p (q %/ p).+1.
+Definition dihedral_gtype := gtype q 2 q.-1.
+Definition semidihedral_gtype := gtype q 2 (q %/ p).-1.
+Definition quaternion_kernel :=
+ <<[set u | u ^+ 2 == 1] :\: [set u ^+ 2 | u in [set: gtype q 4 q.-1]]>>.
+Definition quaternion_gtype :=
+ locked_with gtype_key (coset_groupType quaternion_kernel).
+
+End SpecializeExtremals.
+
+Notation "''Mod_' m" := (modular_gtype m) : type_scope.
+Notation "''Mod_' m" := [set: gsort 'Mod_m] : group_scope.
+Notation "''Mod_' m" := [set: gsort 'Mod_m]%G : Group_scope.
+
+Notation "''D_' m" := (dihedral_gtype m) : type_scope.
+Notation "''D_' m" := [set: gsort 'D_m] : group_scope.
+Notation "''D_' m" := [set: gsort 'D_m]%G : Group_scope.
+
+Notation "''SD_' m" := (semidihedral_gtype m) : type_scope.
+Notation "''SD_' m" := [set: gsort 'SD_m] : group_scope.
+Notation "''SD_' m" := [set: gsort 'SD_m]%G : Group_scope.
+
+Notation "''Q_' m" := (quaternion_gtype m) : type_scope.
+Notation "''Q_' m" := [set: gsort 'Q_m] : group_scope.
+Notation "''Q_' m" := [set: gsort 'Q_m]%G : Group_scope.
+
+Section ExtremalTheory.
+
+Implicit Types (gT : finGroupType) (p q m n : nat).
+
+(* This is Aschbacher (23.3), with the isomorphism made explicit, and a *)
+(* slightly reworked case analysis on the prime and exponent; in particular *)
+(* the inverting involution is available for all non-trivial p-cycles. *)
+Lemma cyclic_pgroup_Aut_structure gT p (G : {group gT}) :
+ p.-group G -> cyclic G -> G :!=: 1 ->
+ let q := #|G| in let n := (logn p q).-1 in
+ let A := Aut G in let P := 'O_p(A) in let F := 'O_p^'(A) in
+ exists m : {perm gT} -> 'Z_q,
+ [/\ [/\ {in A & G, forall a x, x ^+ m a = a x},
+ m 1 = 1%R /\ {in A &, {morph m : a b / a * b >-> (a * b)%R}},
+ {in A &, injective m} /\ image m A =i GRing.unit,
+ forall k, {in A, {morph m : a / a ^+ k >-> (a ^+ k)%R}}
+ & {in A, {morph m : a / a^-1 >-> (a^-1)%R}}],
+ [/\ abelian A, cyclic F, #|F| = p.-1
+ & [faithful F, on 'Ohm_1(G) | [Aut G]]]
+ & if n == 0%N then A = F else
+ exists t, [/\ t \in A, #[t] = 2, m t = - 1%R
+ & if odd p then
+ [/\ cyclic A /\ cyclic P,
+ exists s, [/\ s \in A, #[s] = (p ^ n)%N, m s = p.+1%:R & P = <[s]>]
+ & exists s0, [/\ s0 \in A, #[s0] = p, m s0 = (p ^ n).+1%:R
+ & 'Ohm_1(P) = <[s0]>]]
+ else if n == 1%N then A = <[t]>
+ else exists s,
+ [/\ s \in A, #[s] = (2 ^ n.-1)%N, m s = 5%:R, <[s]> \x <[t]> = A
+ & exists s0, [/\ s0 \in A, #[s0] = 2, m s0 = (2 ^ n).+1%:R,
+ m (s0 * t) = (2 ^ n).-1%:R & 'Ohm_1(<[s]>) = <[s0]>]]]].
+Proof.
+move=> pG cycG ntG q n0 A P F; have [p_pr p_dvd_G [n oG]] := pgroup_pdiv pG ntG.
+have [x0 defG] := cyclicP cycG; have Gx0: x0 \in G by rewrite defG cycle_id.
+rewrite {1}/q oG pfactorK //= in n0 *; rewrite {}/n0.
+have [p_gt1 min_p] := primeP p_pr; have p_gt0 := ltnW p_gt1.
+have q_gt1: q > 1 by rewrite cardG_gt1.
+have cAA: abelian A := Aut_cyclic_abelian cycG; have nilA := abelian_nil cAA.
+have oA: #|A| = (p.-1 * p ^ n)%N.
+ by rewrite card_Aut_cyclic // oG totient_pfactor.
+have [sylP hallF]: p.-Sylow(A) P /\ p^'.-Hall(A) F.
+ by rewrite !nilpotent_pcore_Hall.
+have [defPF tiPF]: P * F = A /\ P :&: F = 1.
+ by case/dprodP: (nilpotent_pcoreC p nilA).
+have oP: #|P| = (p ^ n)%N.
+ by rewrite (card_Hall sylP) oA p_part logn_Gauss ?coprimenP ?pfactorK.
+have oF: #|F| = p.-1.
+ apply/eqP; rewrite -(@eqn_pmul2l #|P|) ?cardG_gt0 // -TI_cardMg // defPF.
+ by rewrite oA oP mulnC.
+have [m' [inj_m' defA def_m']]: exists m' : {morphism units_Zp q >-> {perm gT}},
+ [/\ 'injm m', m' @* setT = A & {in G, forall x u, m' u x = x ^+ val u}].
+- rewrite /A /q defG; exists (Zp_unit_morphism x0).
+ by have [->]:= isomP (Zp_unit_isom x0); split=> // y Gy u; rewrite permE Gy.
+pose m (a : {perm gT}) : 'Z_q := val (invm inj_m' a).
+have{def_m'} def_m: {in A & G, forall a x, x ^+ m a = a x}.
+ by move=> a x Aa Gx /=; rewrite -{2}[a](invmK inj_m') ?defA ?def_m'.
+have m1: m 1 = 1%R by rewrite /m morph1.
+have mM: {in A &, {morph m : a b / a * b >-> (a * b)%R}}.
+ by move=> a b Aa Ab; rewrite /m morphM ?defA.
+have mX k: {in A, {morph m : a / a ^+ k >-> (a ^+ k)%R}}.
+ by elim: k => // k IHk a Aa; rewrite expgS exprS mM ?groupX ?IHk.
+have inj_m: {in A &, injective m}.
+ apply: can_in_inj (fun u => m' (insubd (1 : {unit 'Z_q}) u)) _ => a Aa.
+ by rewrite valKd invmK ?defA.
+have{defA} im_m: image m A =i GRing.unit.
+ move=> u; apply/imageP/idP=> [[a Aa ->]| Uu]; first exact: valP.
+ exists (m' (Sub u Uu)) => /=; first by rewrite -defA mem_morphim ?inE.
+ by rewrite /m invmE ?inE.
+have mV: {in A, {morph m : a / a^-1 >-> (a^-1)%R}}.
+ move=> a Aa /=; rewrite -div1r; apply: canRL (mulrK (valP _)) _.
+ by rewrite -mM ?groupV ?mulVg.
+have inv_m (u : 'Z_q) : coprime q u -> {a | a \in A & m a = u}.
+ rewrite -?unitZpE // natr_Zp -im_m => m_u.
+ by exists (iinv m_u); [exact: mem_iinv | rewrite f_iinv].
+have [cycF ffulF]: cyclic F /\ [faithful F, on 'Ohm_1(G) | [Aut G]].
+ have Um0 a: ((m a)%:R : 'F_p) \in GRing.unit.
+ have: m a \in GRing.unit by exact: valP.
+ by rewrite -{1}[m a]natr_Zp unitFpE ?unitZpE // {1}/q oG coprime_pexpl.
+ pose fm0 a := FinRing.unit 'F_p (Um0 a).
+ have natZqp u: (u%:R : 'Z_q)%:R = u %:R :> 'F_p.
+ by rewrite val_Zp_nat // -Fp_nat_mod // modn_dvdm ?Fp_nat_mod.
+ have m0M: {in A &, {morph fm0 : a b / a * b}}.
+ move=> a b Aa Ab; apply: val_inj; rewrite /= -natrM mM //.
+ by rewrite -[(_ * _)%R]Zp_nat natZqp.
+ pose m0 : {morphism A >-> {unit 'F_p}} := Morphism m0M.
+ have im_m0: m0 @* A = [set: {unit 'F_p}].
+ apply/setP=> [[/= u Uu]]; rewrite in_setT morphimEdom; apply/imsetP.
+ have [|a Aa m_a] := inv_m u%:R.
+ by rewrite {1}[q]oG coprime_pexpl // -unitFpE // natZqp natr_Zp.
+ by exists a => //; apply: val_inj; rewrite /= m_a natZqp natr_Zp.
+ have [x1 defG1]: exists x1, 'Ohm_1(G) = <[x1]>.
+ by apply/cyclicP; exact: cyclicS (Ohm_sub _ _) cycG.
+ have ox1: #[x1] = p by rewrite orderE -defG1 (Ohm1_cyclic_pgroup_prime _ pG).
+ have Gx1: x1 \in G by rewrite -cycle_subG -defG1 Ohm_sub.
+ have ker_m0: 'ker m0 = 'C('Ohm_1(G) | [Aut G]).
+ apply/setP=> a; rewrite inE in_setI; apply: andb_id2l => Aa.
+ rewrite 3!inE /= -2!val_eqE /= val_Fp_nat // [1 %% _]modn_small // defG1.
+ apply/idP/subsetP=> [ma1 x1i | ma1].
+ case/cycleP=> i ->{x1i}; rewrite inE gactX // -[_ a]def_m //.
+ by rewrite -(expg_mod_order x1) ox1 (eqP ma1).
+ have:= ma1 x1 (cycle_id x1); rewrite inE -[_ a]def_m //.
+ by rewrite (eq_expg_mod_order x1 _ 1) ox1 (modn_small p_gt1).
+ have card_units_Fp: #|[set: {unit 'F_p}]| = p.-1.
+ by rewrite card_units_Zp // pdiv_id // (@totient_pfactor p 1) ?muln1.
+ have ker_m0_P: 'ker m0 = P.
+ apply: nilpotent_Hall_pcore nilA _.
+ rewrite pHallE -(card_Hall sylP) oP subsetIl /=.
+ rewrite -(@eqn_pmul2r #|m0 @* A|) ?cardG_gt0 //; apply/eqP.
+ rewrite -{1}(card_isog (first_isog _)) card_quotient ?ker_norm //.
+ by rewrite Lagrange ?subsetIl // oA im_m0 mulnC card_units_Fp.
+ have inj_m0: 'ker_F m0 \subset [1] by rewrite setIC ker_m0_P tiPF.
+ split; last by rewrite /faithful -ker_m0.
+ have isogF: F \isog [set: {unit 'F_p}].
+ have sFA: F \subset A by exact: pcore_sub.
+ apply/isogP; exists (restrm_morphism sFA m0); first by rewrite ker_restrm.
+ apply/eqP; rewrite eqEcard subsetT card_injm ?ker_restrm //= oF.
+ by rewrite card_units_Fp.
+ rewrite (isog_cyclic isogF) pdiv_id // -ox1 (isog_cyclic (Zp_unit_isog x1)).
+ by rewrite Aut_prime_cyclic // -orderE ox1.
+exists m; split=> {im_m mV}//; have [n0 | n_gt0] := posnP n.
+ by apply/eqP; rewrite eq_sym eqEcard pcore_sub oF oA n0 muln1 /=.
+have [t At mt]: {t | t \in A & m t = -1}.
+ apply: inv_m; rewrite /= Zp_cast // coprime_modr modn_small // subn1.
+ by rewrite coprimenP // ltnW.
+have ot: #[t] = 2.
+ apply/eqP; rewrite eqn_leq order_gt1 dvdn_leq ?order_dvdn //=.
+ apply/eqP; move/(congr1 m); apply/eqP; rewrite mt m1 eq_sym -subr_eq0.
+ rewrite opprK -val_eqE /= Zp_cast ?modn_small // /q oG ltnW //.
+ by rewrite (leq_trans (_ : 2 ^ 2 <= p ^ 2)) ?leq_sqr ?leq_exp2l.
+ by apply/eqP; apply: inj_m; rewrite ?groupX ?group1 ?mX // mt -signr_odd.
+exists t; split=> //.
+case G4: (~~ odd p && (n == 1%N)).
+ case: (even_prime p_pr) G4 => [p2 | -> //]; rewrite p2 /=; move/eqP=> n1.
+ rewrite n1 /=; apply/eqP; rewrite eq_sym eqEcard cycle_subG At /=.
+ by rewrite -orderE oA ot p2 n1.
+pose e0 : nat := ~~ odd p.
+have{inv_m} [s As ms]: {s | s \in A & m s = (p ^ e0.+1).+1%:R}.
+ apply: inv_m; rewrite val_Zp_nat // coprime_modr /q oG coprime_pexpl //.
+ by rewrite -(@coprime_pexpl e0.+1) // coprimenS.
+have lt_e0_n: e0 < n.
+ by rewrite /e0; case: (~~ _) G4 => //=; rewrite ltn_neqAle eq_sym => ->.
+pose s0 := s ^+ (p ^ (n - e0.+1)).
+have [ms0 os0]: m s0 = (p ^ n).+1%:R /\ #[s0] = p.
+ have m_se e:
+ exists2 k, k = 1 %[mod p] & m (s ^+ (p ^ e)) = (k * p ^ (e + e0.+1)).+1%:R.
+ - elim: e => [|e [k k1 IHe]]; first by exists 1%N; rewrite ?mul1n.
+ rewrite expnSr expgM mX ?groupX // {}IHe -natrX -(add1n (k * _)).
+ rewrite expnDn -(prednK p_gt0) 2!big_ord_recl /= prednK // !exp1n bin1.
+ rewrite bin0 muln1 mul1n mulnCA -expnS (addSn e).
+ set f := (e + _)%N; set sum := (\sum_i _)%N.
+ exists (sum %/ p ^ f.+2 * p + k)%N; first by rewrite modnMDl.
+ rewrite -(addnC k) mulnDl -mulnA -expnS divnK // {}/sum.
+ apply big_ind => [||[i _] /= _]; [exact: dvdn0 | exact: dvdn_add |].
+ rewrite exp1n mul1n /bump !add1n expnMn mulnCA dvdn_mull // -expnM.
+ case: (ltnP f.+1 (f * i.+2)) => [le_f_fi|].
+ by rewrite dvdn_mull ?dvdn_exp2l.
+ rewrite {1}mulnS -(addn1 f) leq_add2l {}/f addnS /e0.
+ case: i e => [] // [] //; case odd_p: (odd p) => //= _.
+ by rewrite bin2odd // mulnAC dvdn_mulr.
+ have [[|d]] := m_se (n - e0.+1)%N; first by rewrite mod0n modn_small.
+ move/eqP; rewrite -/s0 eqn_mod_dvd ?subn1 //=; case/dvdnP=> f -> {d}.
+ rewrite subnK // mulSn -mulnA -expnS -addSn natrD natrM -oG char_Zp //.
+ rewrite mulr0 addr0 => m_s0; split => //.
+ have [d _] := m_se (n - e0)%N; rewrite -subnSK // expnSr expgM -/s0.
+ rewrite addSn subnK // -oG mulrS natrM char_Zp // {d}mulr0 addr0.
+ move/eqP; rewrite -m1 (inj_in_eq inj_m) ?group1 ?groupX // -order_dvdn.
+ move/min_p; rewrite order_eq1; case/predU1P=> [s0_1 | ]; last by move/eqP.
+ move/eqP: m_s0; rewrite eq_sym s0_1 m1 -subr_eq0 mulrSr addrK -val_eqE /=.
+ have pf_gt0: p ^ _ > 0 by move=> e; rewrite expn_gt0 p_gt0.
+ by rewrite val_Zp_nat // /q oG [_ == _]pfactor_dvdn // pfactorK ?ltnn.
+have os: #[s] = (p ^ (n - e0))%N.
+ have: #[s] %| p ^ (n - e0).
+ by rewrite order_dvdn -subnSK // expnSr expgM -order_dvdn os0.
+ case/dvdn_pfactor=> // d; rewrite leq_eqVlt.
+ case/predU1P=> [-> // | lt_d os]; case/idPn: (p_gt1); rewrite -os0.
+ by rewrite order_gt1 negbK -order_dvdn os dvdn_exp2l // -ltnS -subSn.
+have p_s: p.-elt s by rewrite /p_elt os pnat_exp ?pnat_id.
+have defS1: 'Ohm_1(<[s]>) = <[s0]>.
+ apply/eqP; rewrite eq_sym eqEcard cycle_subG -orderE os0.
+ rewrite (Ohm1_cyclic_pgroup_prime _ p_s) ?cycle_cyclic ?leqnn ?cycle_eq1 //=.
+ rewrite (OhmE _ p_s) mem_gen ?groupX //= !inE mem_cycle //.
+ by rewrite -order_dvdn os0 ?dvdnn.
+ by apply/eqP=> s1; rewrite -os0 /s0 s1 expg1n order1 in p_gt1.
+case: (even_prime p_pr) => [p2 | oddp]; last first.
+ rewrite {+}/e0 oddp subn0 in s0 os0 ms0 os ms defS1 *.
+ have [f defF] := cyclicP cycF; have defP: P = <[s]>.
+ apply/eqP; rewrite eq_sym eqEcard -orderE oP os leqnn andbT.
+ by rewrite cycle_subG (mem_normal_Hall sylP) ?pcore_normal.
+ rewrite defP; split; last 1 [by exists s | by exists s0; rewrite ?groupX].
+ rewrite -defPF defP defF -cycleM ?cycle_cyclic // /order.
+ by red; rewrite (centsP cAA) // -cycle_subG -defF pcore_sub.
+ by rewrite -defF -defP (pnat_coprime (pcore_pgroup _ _) (pcore_pgroup _ _)).
+rewrite {+}/e0 p2 subn1 /= in s0 os0 ms0 os ms G4 defS1 lt_e0_n *.
+rewrite G4; exists s; split=> //; last first.
+ exists s0; split; rewrite ?groupX //; apply/eqP; rewrite mM ?groupX //.
+ rewrite ms0 mt eq_sym mulrN1 -subr_eq0 opprK -natrD -addSnnS.
+ by rewrite prednK ?expn_gt0 // addnn -mul2n -expnS -p2 -oG char_Zp.
+suffices TIst: <[s]> :&: <[t]> = 1.
+ rewrite dprodE //; last by rewrite (sub_abelian_cent2 cAA) ?cycle_subG.
+ apply/eqP; rewrite eqEcard mulG_subG !cycle_subG As At oA.
+ by rewrite TI_cardMg // -!orderE os ot p2 mul1n /= -expnSr prednK.
+rewrite setIC; apply: prime_TIg; first by rewrite -orderE ot.
+rewrite cycle_subG; apply/negP=> St.
+have: t \in <[s0]>.
+ by rewrite -defS1 (OhmE _ p_s) mem_gen // !inE St -order_dvdn ot p2.
+have ->: <[s0]> = [set 1; s0].
+ apply/eqP; rewrite eq_sym eqEcard subUset !sub1set group1 cycle_id /=.
+ by rewrite -orderE cards2 eq_sym -order_gt1 os0.
+rewrite !inE -order_eq1 ot /=; move/eqP; move/(congr1 m); move/eqP.
+rewrite mt ms0 eq_sym -subr_eq0 opprK -mulrSr.
+rewrite -val_eqE [val _]val_Zp_nat //= /q oG p2 modn_small //.
+by rewrite -addn3 expnS mul2n -addnn leq_add2l (ltn_exp2l 1).
+Qed.
+
+Definition extremal_generators gT (A : {set gT}) p n xy :=
+ let: (x, y) := xy in
+ [/\ #|A| = (p ^ n)%N, x \in A, #[x] = (p ^ n.-1)%N & y \in A :\: <[x]>].
+
+Lemma extremal_generators_facts gT (G : {group gT}) p n x y :
+ prime p -> extremal_generators G p n (x, y) ->
+ [/\ p.-group G, maximal <[x]> G, <[x]> <| G,
+ <[x]> * <[y]> = G & <[y]> \subset 'N(<[x]>)].
+Proof.
+move=> p_pr [oG Gx ox] /setDP[Gy notXy].
+have pG: p.-group G by rewrite /pgroup oG pnat_exp pnat_id.
+have maxX: maximal <[x]> G.
+ rewrite p_index_maximal -?divgS ?cycle_subG // -orderE oG ox.
+ case: (n) oG => [|n' _]; last by rewrite -expnB ?subSnn ?leqnSn ?prime_gt0.
+ move/eqP; rewrite -trivg_card1; case/trivgPn.
+ by exists y; rewrite // (group1_contra notXy).
+have nsXG := p_maximal_normal pG maxX; split=> //.
+ by apply: mulg_normal_maximal; rewrite ?cycle_subG.
+by rewrite cycle_subG (subsetP (normal_norm nsXG)).
+Qed.
+
+Section ModularGroup.
+
+Variables p n : nat.
+Let m := (p ^ n)%N.
+Let q := (p ^ n.-1)%N.
+Let r := (p ^ n.-2)%N.
+
+Hypotheses (p_pr : prime p) (n_gt2 : n > 2).
+Let p_gt1 := prime_gt1 p_pr.
+Let p_gt0 := ltnW p_gt1.
+Let def_n := esym (subnKC n_gt2).
+Let def_p : pdiv m = p. Proof. by rewrite /m def_n pdiv_pfactor. Qed.
+Let def_q : m %/ p = q. Proof. by rewrite /m /q def_n expnS mulKn. Qed.
+Let def_r : q %/ p = r. Proof. by rewrite /r /q def_n expnS mulKn. Qed.
+Let ltqm : q < m. Proof. by rewrite ltn_exp2l // def_n. Qed.
+Let ltrq : r < q. Proof. by rewrite ltn_exp2l // def_n. Qed.
+Let r_gt0 : 0 < r. Proof. by rewrite expn_gt0 ?p_gt0. Qed.
+Let q_gt1 : q > 1. Proof. exact: leq_ltn_trans r_gt0 ltrq. Qed.
+
+Lemma card_modular_group : #|'Mod_(p ^ n)| = (p ^ n)%N.
+Proof. by rewrite Extremal.card def_p ?def_q // -expnS def_n. Qed.
+
+Lemma Grp_modular_group :
+ 'Mod_(p ^ n) \isog Grp (x : y : (x ^+ q, y ^+ p, x ^ y = x ^+ r.+1)).
+Proof.
+rewrite /modular_gtype def_p def_q def_r; apply: Extremal.Grp => //.
+set B := <[_]>; have Bb: Zp1 \in B by exact: cycle_id.
+have oB: #|B| = q by rewrite -orderE order_Zp1 Zp_cast.
+have cycB: cyclic B by rewrite cycle_cyclic.
+have pB: p.-group B by rewrite /pgroup oB pnat_exp ?pnat_id.
+have ntB: B != 1 by rewrite -cardG_gt1 oB.
+have [] := cyclic_pgroup_Aut_structure pB cycB ntB.
+rewrite oB pfactorK //= -/B -(expg_znat r.+1 Bb) oB => mB [[def_mB _ _ _ _] _].
+rewrite {1}def_n /= => [[t [At ot mBt]]].
+have [p2 | ->] := even_prime p_pr; last first.
+ by case=> _ _ [s [As os mBs _]]; exists s; rewrite os -mBs def_mB.
+rewrite {1}p2 /= -2!eqSS -addn2 -2!{1}subn1 -subnDA subnK 1?ltnW //.
+case: eqP => [n3 _ | _ [_ [_ _ _ _ [s [As os mBs _ _]{t At ot mBt}]]]].
+ by exists t; rewrite At ot -def_mB // mBt /q /r p2 n3.
+by exists s; rewrite As os -def_mB // mBs /r p2.
+Qed.
+
+Definition modular_group_generators gT (xy : gT * gT) :=
+ let: (x, y) := xy in #[y] = p /\ x ^ y = x ^+ r.+1.
+
+Lemma generators_modular_group gT (G : {group gT}) :
+ G \isog 'Mod_m ->
+ exists2 xy, extremal_generators G p n xy & modular_group_generators xy.
+Proof.
+case/(isoGrpP _ Grp_modular_group); rewrite card_modular_group // -/m => oG.
+case/existsP=> -[x y] /= /eqP[defG xq yp xy].
+rewrite norm_joinEr ?norms_cycle ?xy ?mem_cycle // in defG.
+have [Gx Gy]: x \in G /\ y \in G.
+ by apply/andP; rewrite -!cycle_subG -mulG_subG defG.
+have notXy: y \notin <[x]>.
+ apply: contraL ltqm; rewrite -cycle_subG -oG -defG; move/mulGidPl->.
+ by rewrite -leqNgt dvdn_leq ?(ltnW q_gt1) // order_dvdn xq.
+have oy: #[y] = p by exact: nt_prime_order (group1_contra notXy).
+exists (x, y) => //=; split; rewrite ?inE ?notXy //.
+apply/eqP; rewrite -(eqn_pmul2r p_gt0) -expnSr -{1}oy (ltn_predK n_gt2) -/m.
+by rewrite -TI_cardMg ?defG ?oG // setIC prime_TIg ?cycle_subG // -orderE oy.
+Qed.
+
+(* This is an adaptation of Aschbacher, exercise 8.2: *)
+(* - We allow an alternative to the #[x] = p ^ n.-1 condition that meshes *)
+(* better with the modular_Grp lemma above. *)
+(* - We state explicitly some "obvious" properties of G, namely that G is *)
+(* the non-abelian semi-direct product <[x]> ><| <[y]> and that y ^+ j *)
+(* acts on <[x]> via z |-> z ^+ (j * p ^ n.-2).+1 *)
+(* - We also give the values of the 'Mho^k(G). *)
+(* - We corrected a pair of typos. *)
+Lemma modular_group_structure gT (G : {group gT}) x y :
+ extremal_generators G p n (x, y) ->
+ G \isog 'Mod_m -> modular_group_generators (x, y) ->
+ let X := <[x]> in
+ [/\ [/\ X ><| <[y]> = G, ~~ abelian G
+ & {in X, forall z j, z ^ (y ^+ j) = z ^+ (j * r).+1}],
+ [/\ 'Z(G) = <[x ^+ p]>, 'Phi(G) = 'Z(G) & #|'Z(G)| = r],
+ [/\ G^`(1) = <[x ^+ r]>, #|G^`(1)| = p & nil_class G = 2],
+ forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (p ^ k)]>
+ & if (p, n) == (2, 3) then 'Ohm_1(G) = G else
+ forall k, 0 < k < n.-1 ->
+ <[x ^+ (p ^ (n - k.+1))]> \x <[y]> = 'Ohm_k(G)
+ /\ #|'Ohm_k(G)| = (p ^ k.+1)%N].
+Proof.
+move=> genG isoG [oy xy] X.
+have [oG Gx ox /setDP[Gy notXy]] := genG; rewrite -/m -/q in ox oG.
+have [pG _ nsXG defXY nXY] := extremal_generators_facts p_pr genG.
+have [sXG nXG] := andP nsXG; have sYG: <[y]> \subset G by rewrite cycle_subG.
+have n1_gt1: n.-1 > 1 by [rewrite def_n]; have n1_gt0 := ltnW n1_gt1.
+have def_n1 := prednK n1_gt0.
+have def_m: (q * p)%N = m by rewrite -expnSr /m def_n.
+have notcxy: y \notin 'C[x].
+ apply: contraL (introT eqP xy); move/cent1P=> cxy.
+ rewrite /conjg -cxy // eq_mulVg1 expgS !mulKg -order_dvdn ox.
+ by rewrite pfactor_dvdn ?expn_gt0 ?p_gt0 // pfactorK // -ltnNge prednK.
+have tiXY: <[x]> :&: <[y]> = 1.
+ rewrite setIC prime_TIg -?orderE ?oy //; apply: contra notcxy.
+ by rewrite cycle_subG; apply: subsetP; rewrite cycle_subG cent1id.
+have notcGG: ~~ abelian G.
+ by rewrite -defXY abelianM !cycle_abelian cent_cycle cycle_subG.
+have cXpY: <[y]> \subset 'C(<[x ^+ p]>).
+ rewrite cent_cycle cycle_subG cent1C (sameP cent1P commgP) /commg conjXg xy.
+ by rewrite -expgM mulSn expgD mulKg -expnSr def_n1 -/q -ox expg_order.
+have oxp: #[x ^+ p] = r by rewrite orderXdiv ox ?dvdn_exp //.
+have [sZG nZG] := andP (center_normal G).
+have defZ: 'Z(G) = <[x ^+ p]>.
+ apply/eqP; rewrite eq_sym eqEcard subsetI -{2}defXY centM subsetI cent_cycle.
+ rewrite 2!cycle_subG !groupX ?cent1id //= centsC cXpY /= -orderE oxp leqNgt.
+ apply: contra notcGG => gtZr; apply: cyclic_center_factor_abelian.
+ rewrite (dvdn_prime_cyclic p_pr) // card_quotient //.
+ rewrite -(dvdn_pmul2l (cardG_gt0 'Z(G))) Lagrange // oG -def_m dvdn_pmul2r //.
+ case/p_natP: (pgroupS sZG pG) gtZr => k ->.
+ by rewrite ltn_exp2l // def_n1; exact: dvdn_exp2l.
+have Zxr: x ^+ r \in 'Z(G) by rewrite /r def_n expnS expgM defZ mem_cycle.
+have rxy: [~ x, y] = x ^+ r by rewrite /commg xy expgS mulKg.
+have defG': G^`(1) = <[x ^+ r]>.
+ case/setIP: Zxr => _; rewrite -rxy -defXY -(norm_joinEr nXY).
+ exact: der1_joing_cycles.
+have oG': #|G^`(1)| = p.
+ by rewrite defG' -orderE orderXdiv ox /q -def_n1 ?dvdn_exp2l // expnS mulnK.
+have sG'Z: G^`(1) \subset 'Z(G) by rewrite defG' cycle_subG.
+have nil2_G: nil_class G = 2.
+ by apply/eqP; rewrite eqn_leq andbC ltnNge nil_class1 notcGG nil_class2.
+have XYp: {in X & <[y]>, forall z t,
+ (z * t) ^+ p \in z ^+ p *: <[x ^+ r ^+ 'C(p, 2)]>}.
+- move=> z t Xz Yt; have Gz := subsetP sXG z Xz; have Gt := subsetP sYG t Yt.
+ have Rtz: [~ t, z] \in G^`(1) by exact: mem_commg.
+ have cGtz: [~ t, z] \in 'C(G) by case/setIP: (subsetP sG'Z _ Rtz).
+ rewrite expMg_Rmul /commute ?(centP cGtz) //.
+ have ->: t ^+ p = 1 by apply/eqP; rewrite -order_dvdn -oy order_dvdG.
+ rewrite defG' in Rtz; case/cycleP: Rtz => i ->.
+ by rewrite mem_lcoset mulg1 mulKg expgAC mem_cycle.
+have defMho: 'Mho^1(G) = <[x ^+ p]>.
+ apply/eqP; rewrite eqEsubset cycle_subG (Mho_p_elt 1) ?(mem_p_elt pG) //.
+ rewrite andbT (MhoE 1 pG) gen_subG -defXY; apply/subsetP=> ztp.
+ case/imsetP=> zt; case/imset2P=> z t Xz Yt -> -> {zt ztp}.
+ apply: subsetP (XYp z t Xz Yt); case/cycleP: Xz => i ->.
+ by rewrite expgAC mul_subG ?sub1set ?mem_cycle //= -defZ cycle_subG groupX.
+split=> //; try exact: extend_cyclic_Mho.
+- rewrite sdprodE //; split=> // z; case/cycleP=> i ->{z} j.
+ rewrite conjXg -expgM mulnC expgM actX; congr (_ ^+ i).
+ elim: j {i} => //= j ->; rewrite conjXg xy -!expgM mulnS mulSn addSn.
+ rewrite addnA -mulSn -addSn expgD mulnCA (mulnC j).
+ rewrite {3}/r def_n expnS mulnA -expnSr def_n1 -/q -ox -mulnA expgM.
+ by rewrite expg_order expg1n mulg1.
+- by rewrite (Phi_joing pG) defMho -defZ (joing_idPr _) ?defZ.
+have G1y: y \in 'Ohm_1(G).
+ by rewrite (OhmE _ pG) mem_gen // !inE Gy -order_dvdn oy /=.
+case: eqP => [[p2 n3] | notG8 k]; last case/andP=> k_gt0 lt_k_n1.
+ apply/eqP; rewrite eqEsubset Ohm_sub -{1}defXY mulG_subG !cycle_subG.
+ rewrite G1y -(groupMr _ G1y) /= (OhmE _ pG) mem_gen // !inE groupM //.
+ rewrite /q /r p2 n3 in oy ox xy *.
+ by rewrite expgS -mulgA -{1}(invg2id oy) -conjgE xy -expgS -order_dvdn ox.
+have le_k_n2: k <= n.-2 by rewrite -def_n1 in lt_k_n1.
+suffices{lt_k_n1} defGk: <[x ^+ (p ^ (n - k.+1))]> \x <[y]> = 'Ohm_k(G).
+ split=> //; case/dprodP: defGk => _ <- _ tiXkY; rewrite expnSr TI_cardMg //.
+ rewrite -!orderE oy (subnDA 1) subn1 orderXdiv ox ?dvdn_exp2l ?leq_subr //.
+ by rewrite /q -{1}(subnK (ltnW lt_k_n1)) expnD mulKn // expn_gt0 p_gt0.
+suffices{k k_gt0 le_k_n2} defGn2: <[x ^+ p]> \x <[y]> = 'Ohm_(n.-2)(G).
+ have:= Ohm_dprod k defGn2; have p_xp := mem_p_elt pG (groupX p Gx).
+ rewrite (Ohm_p_cycle _ p_xp) (Ohm_p_cycle _ (mem_p_elt pG Gy)) oxp oy.
+ rewrite pfactorK ?(pfactorK 1) // (eqnP k_gt0) expg1 -expgM -expnS.
+ rewrite -subSn // -subSS def_n1 def_n => -> /=; rewrite subnSK // subn2.
+ by apply/eqP; rewrite eqEsubset OhmS ?Ohm_sub //= -{1}Ohm_id OhmS ?Ohm_leq.
+rewrite dprodEY //=; last by apply/trivgP; rewrite -tiXY setSI ?cycleX.
+apply/eqP; rewrite eqEsubset join_subG !cycle_subG /= {-2}(OhmE _ pG) -/r.
+rewrite def_n (subsetP (Ohm_leq G (ltn0Sn _))) // mem_gen /=; last first.
+ by rewrite !inE -order_dvdn oxp groupX /=.
+rewrite gen_subG /= cent_joinEr // -defXY; apply/subsetP=> uv; case/setIP.
+case/imset2P=> u v Xu Yv ->{uv}; rewrite /r inE def_n expnS expgM.
+case/lcosetP: (XYp u v Xu Yv) => _ /cycleP[j ->] ->.
+case/cycleP: Xu => i ->{u}; rewrite -!(expgM, expgD) -order_dvdn ox.
+rewrite (mulnC r) /r {1}def_n expnSr mulnA -mulnDl -mulnA -expnS.
+rewrite subnSK // subn2 /q -def_n1 expnS dvdn_pmul2r // dvdn_addl.
+ by case/dvdnP=> k ->; rewrite mulnC expgM mem_mulg ?mem_cycle.
+case: (ltngtP n 3) => [|n_gt3|n3]; first by rewrite ltnNge n_gt2.
+ by rewrite -subnSK // expnSr mulnA dvdn_mull.
+case: (even_prime p_pr) notG8 => [-> | oddp _]; first by rewrite n3.
+by rewrite bin2odd // -!mulnA dvdn_mulr.
+Qed.
+
+End ModularGroup.
+
+(* Basic properties of dihedral groups; these will be refined for dihedral *)
+(* 2-groups in the section on extremal 2-groups. *)
+Section DihedralGroup.
+
+Variable q : nat.
+Hypothesis q_gt1 : q > 1.
+Let m := q.*2.
+
+Let def2 : pdiv m = 2.
+Proof.
+apply/eqP; rewrite /m -mul2n eqn_leq pdiv_min_dvd ?dvdn_mulr //.
+by rewrite prime_gt1 // pdiv_prime // (@leq_pmul2l 2 1) ltnW.
+Qed.
+
+Let def_q : m %/ pdiv m = q. Proof. by rewrite def2 divn2 half_double. Qed.
+
+Section Dihedral_extension.
+
+Variable p : nat.
+Hypotheses (p_gt1 : p > 1) (even_p : 2 %| p).
+Local Notation ED := [set: gsort (Extremal.gtype q p q.-1)].
+
+Lemma card_ext_dihedral : #|ED| = (p./2 * m)%N.
+Proof. by rewrite Extremal.card // /m -mul2n -divn2 mulnA divnK. Qed.
+
+Lemma Grp_ext_dihedral : ED \isog Grp (x : y : (x ^+ q, y ^+ p, x ^ y = x^-1)).
+Proof.
+suffices isoED: ED \isog Grp (x : y : (x ^+ q, y ^+ p, x ^ y = x ^+ q.-1)).
+ move=> gT G; rewrite isoED.
+ apply: eq_existsb => [[x y]] /=; rewrite !xpair_eqE.
+ congr (_ && _); apply: andb_id2l; move/eqP=> xq1; congr (_ && (_ == _)).
+ by apply/eqP; rewrite eq_sym eq_invg_mul -expgS (ltn_predK q_gt1) xq1.
+have unitrN1 : - 1 \in GRing.unit by move=> R; rewrite unitrN unitr1.
+pose uN1 := FinRing.unit ('Z_#[Zp1 : 'Z_q]) (unitrN1 _).
+apply: Extremal.Grp => //; exists (Zp_unitm uN1).
+rewrite Aut_aut order_injm ?injm_Zp_unitm ?in_setT //; split=> //.
+ by rewrite (dvdn_trans _ even_p) // order_dvdn -val_eqE /= mulrNN.
+apply/eqP; rewrite autE ?cycle_id // eq_expg_mod_order /=.
+by rewrite order_Zp1 !Zp_cast // !modn_mod (modn_small q_gt1) subn1.
+Qed.
+
+End Dihedral_extension.
+
+Lemma card_dihedral : #|'D_m| = m.
+Proof. by rewrite /('D_m)%type def_q card_ext_dihedral ?mul1n. Qed.
+
+Lemma Grp_dihedral : 'D_m \isog Grp (x : y : (x ^+ q, y ^+ 2, x ^ y = x^-1)).
+Proof. by rewrite /('D_m)%type def_q; exact: Grp_ext_dihedral. Qed.
+
+Lemma Grp'_dihedral : 'D_m \isog Grp (x : y : (x ^+ 2, y ^+ 2, (x * y) ^+ q)).
+Proof.
+move=> gT G; rewrite Grp_dihedral; apply/existsP/existsP=> [] [[x y]] /=.
+ case/eqP=> <- xq1 y2 xy; exists (x * y, y); rewrite !xpair_eqE /= eqEsubset.
+ rewrite !join_subG !joing_subr !cycle_subG -{3}(mulgK y x) /=.
+ rewrite 2?groupM ?groupV ?mem_gen ?inE ?cycle_id ?orbT //= -mulgA expgS.
+ by rewrite {1}(conjgC x) xy -mulgA mulKg -(expgS y 1) y2 mulg1 xq1 !eqxx.
+case/eqP=> <- x2 y2 xyq; exists (x * y, y); rewrite !xpair_eqE /= eqEsubset.
+rewrite !join_subG !joing_subr !cycle_subG -{3}(mulgK y x) /=.
+rewrite 2?groupM ?groupV ?mem_gen ?inE ?cycle_id ?orbT //= xyq y2 !eqxx /=.
+by rewrite eq_sym eq_invg_mul !mulgA mulgK -mulgA -!(expgS _ 1) x2 y2 mulg1.
+Qed.
+
+End DihedralGroup.
+
+Lemma involutions_gen_dihedral gT (x y : gT) :
+ let G := <<[set x; y]>> in
+ #[x] = 2 -> #[y] = 2 -> x != y -> G \isog 'D_#|G|.
+Proof.
+move=> G ox oy ne_x_y; pose q := #[x * y].
+have q_gt1: q > 1 by rewrite order_gt1 -eq_invg_mul invg_expg ox.
+have homG: G \homg 'D_q.*2.
+ rewrite Grp'_dihedral //; apply/existsP; exists (x, y); rewrite /= !xpair_eqE.
+ by rewrite joing_idl joing_idr -{1}ox -oy !expg_order !eqxx.
+suff oG: #|G| = q.*2 by rewrite oG isogEcard oG card_dihedral ?leqnn ?andbT.
+have: #|G| %| q.*2 by rewrite -card_dihedral ?card_homg.
+have Gxy: <[x * y]> \subset G.
+ by rewrite cycle_subG groupM ?mem_gen ?set21 ?set22.
+have[k oG]: exists k, #|G| = (k * q)%N by apply/dvdnP; rewrite cardSg.
+rewrite oG -mul2n dvdn_pmul2r ?order_gt0 ?dvdn_divisors // !inE /=.
+case/pred2P=> [k1 | -> //]; case/negP: ne_x_y.
+have cycG: cyclic G.
+ apply/cyclicP; exists (x * y); apply/eqP.
+ by rewrite eq_sym eqEcard Gxy oG k1 mul1n leqnn.
+have: <[x]> == <[y]>.
+ by rewrite (eq_subG_cyclic cycG) ?genS ?subsetUl ?subsetUr -?orderE ?ox ?oy.
+by rewrite eqEcard cycle_subG /= cycle2g // !inE -order_eq1 ox; case/andP.
+Qed.
+
+Lemma Grp_2dihedral n : n > 1 ->
+ 'D_(2 ^ n) \isog Grp (x : y : (x ^+ (2 ^ n.-1), y ^+ 2, x ^ y = x^-1)).
+Proof.
+move=> n_gt1; rewrite -(ltn_predK n_gt1) expnS mul2n /=.
+by apply: Grp_dihedral; rewrite (ltn_exp2l 0) // -(subnKC n_gt1).
+Qed.
+
+Lemma card_2dihedral n : n > 1 -> #|'D_(2 ^ n)| = (2 ^ n)%N.
+Proof.
+move=> n_gt1; rewrite -(ltn_predK n_gt1) expnS mul2n /= card_dihedral //.
+by rewrite (ltn_exp2l 0) // -(subnKC n_gt1).
+Qed.
+
+Lemma card_semidihedral n : n > 3 -> #|'SD_(2 ^ n)| = (2 ^ n)%N.
+Proof.
+move=> n_gt3.
+rewrite /('SD__)%type -(subnKC (ltnW (ltnW n_gt3))) pdiv_pfactor //.
+by rewrite // !expnS !mulKn -?expnS ?Extremal.card //= (ltn_exp2l 0).
+Qed.
+
+Lemma Grp_semidihedral n : n > 3 ->
+ 'SD_(2 ^ n) \isog
+ Grp (x : y : (x ^+ (2 ^ n.-1), y ^+ 2, x ^ y = x ^+ (2 ^ n.-2).-1)).
+Proof.
+move=> n_gt3.
+rewrite /('SD__)%type -(subnKC (ltnW (ltnW n_gt3))) pdiv_pfactor //.
+rewrite !expnS !mulKn // -!expnS /=; set q := (2 ^ _)%N.
+have q_gt1: q > 1 by rewrite (ltn_exp2l 0).
+apply: Extremal.Grp => //; set B := <[_]>.
+have oB: #|B| = q by rewrite -orderE order_Zp1 Zp_cast.
+have pB: 2.-group B by rewrite /pgroup oB pnat_exp.
+have ntB: B != 1 by rewrite -cardG_gt1 oB.
+have [] := cyclic_pgroup_Aut_structure pB (cycle_cyclic _) ntB.
+rewrite oB /= pfactorK //= -/B => m [[def_m _ _ _ _] _].
+rewrite -{1 2}(subnKC n_gt3) => [[t [At ot _ [s [_ _ _ defA]]]]].
+case/dprodP: defA => _ defA cst _.
+have{cst defA} cAt: t \in 'C(Aut B).
+ rewrite -defA centM inE -sub_cent1 -cent_cycle centsC cst /=.
+ by rewrite cent_cycle cent1id.
+case=> s0 [As0 os0 _ def_s0t _]; exists (s0 * t).
+rewrite -def_m ?groupM ?cycle_id // def_s0t !Zp_expg !mul1n valZpK Zp_nat.
+rewrite order_dvdn expgMn /commute 1?(centP cAt) // -{1}os0 -{1}ot.
+by rewrite !expg_order mul1g.
+Qed.
+
+Section Quaternion.
+
+Variable n : nat.
+Hypothesis n_gt2 : n > 2.
+Let m := (2 ^ n)%N.
+Let q := (2 ^ n.-1)%N.
+Let r := (2 ^ n.-2)%N.
+Let GrpQ := 'Q_m \isog Grp (x : y : (x ^+ q, y ^+ 2 = x ^+ r, x ^ y = x^-1)).
+Let defQ : #|'Q_m| = m /\ GrpQ.
+Proof.
+have q_gt1 : q > 1 by rewrite (ltn_exp2l 0) // -(subnKC n_gt2).
+have def_m : (2 * q)%N = m by rewrite -expnS (ltn_predK n_gt2).
+have def_q : m %/ pdiv m = q
+ by rewrite /m -(ltn_predK n_gt2) pdiv_pfactor // expnS mulKn.
+have r_gt1 : r > 1 by rewrite (ltn_exp2l 0) // -(subnKC n_gt2).
+have def2r : (2 * r)%N = q by rewrite -expnS /q -(subnKC n_gt2).
+rewrite /GrpQ [@quaternion_gtype _]unlock /quaternion_kernel {}def_q.
+set B := [set: _]; have: B \homg Grp (u : v : (u ^+ q, v ^+ 4, u ^ v = u^-1)).
+ by rewrite -Grp_ext_dihedral ?homg_refl.
+have: #|B| = (q * 4)%N by rewrite card_ext_dihedral // mulnC -muln2 -mulnA.
+rewrite {}/B; move: (Extremal.gtype q 4 _) => gT.
+set B := [set: gT] => oB; set K := _ :\: _.
+case/existsP=> -[u v] /= /eqP[defB uq v4 uv].
+have nUV: <[v]> \subset 'N(<[u]>) by rewrite norms_cycle uv groupV cycle_id.
+rewrite norm_joinEr // in defB.
+have le_ou: #[u] <= q by rewrite dvdn_leq ?expn_gt0 // order_dvdn uq.
+have le_ov: #[v] <= 4 by rewrite dvdn_leq // order_dvdn v4.
+have tiUV: <[u]> :&: <[v]> = 1 by rewrite cardMg_TI // defB oB leq_mul.
+have{le_ou le_ov} [ou ov]: #[u] = q /\ #[v] = 4.
+ have:= esym (leqif_mul (leqif_eq le_ou) (leqif_eq le_ov)).2.
+ by rewrite -TI_cardMg // defB -oB eqxx eqn0Ngt cardG_gt0; do 2!case: eqP=> //.
+have sdB: <[u]> ><| <[v]> = B by rewrite sdprodE.
+have uvj j: u ^ (v ^+ j) = (if odd j then u^-1 else u).
+ elim: j => [|j IHj]; first by rewrite conjg1.
+ by rewrite expgS conjgM uv conjVg IHj (fun_if invg) invgK if_neg.
+have sqrB i j: (u ^+ i * v ^+ j) ^+ 2 = (if odd j then v ^+ 2 else u ^+ i.*2).
+ rewrite expgS; case: ifP => odd_j.
+ rewrite {1}(conjgC (u ^+ i)) conjXg uvj odd_j expgVn -mulgA mulKg.
+ rewrite -expgD addnn -(odd_double_half j) odd_j doubleD addnC /=.
+ by rewrite -(expg_mod _ v4) -!muln2 -mulnA modnMDl.
+ rewrite {2}(conjgC (u ^+ i)) conjXg uvj odd_j mulgA -(mulgA (u ^+ i)).
+ rewrite -expgD addnn -(odd_double_half j) odd_j -2!mul2n mulnA.
+ by rewrite expgM v4 expg1n mulg1 -expgD addnn.
+pose w := u ^+ r * v ^+ 2.
+have Kw: w \in K.
+ rewrite !inE sqrB /= -mul2n def2r uq eqxx andbT -defB.
+ apply/imsetP=> [[_]] /imset2P[_ _ /cycleP[i ->] /cycleP[j ->] ->].
+ apply/eqP; rewrite sqrB; case: ifP => _.
+ rewrite eq_mulgV1 mulgK -order_dvdn ou pfactor_dvdn ?expn_gt0 ?pfactorK //.
+ by rewrite -ltnNge -(subnKC n_gt2).
+ rewrite (canF_eq (mulKg _)); apply/eqP=> def_v2.
+ suffices: v ^+ 2 \in <[u]> :&: <[v]> by rewrite tiUV inE -order_dvdn ov.
+ by rewrite inE {1}def_v2 groupM ?groupV !mem_cycle.
+have ow: #[w] = 2.
+ case/setDP: Kw; rewrite inE -order_dvdn dvdn_divisors // !inE /= order_eq1.
+ by case/orP=> /eqP-> // /imsetP[]; exists 1; rewrite ?inE ?expg1n.
+have defK: K = [set w].
+ apply/eqP; rewrite eqEsubset sub1set Kw andbT subDset setUC.
+ apply/subsetP=> uivj; have: uivj \in B by rewrite inE.
+ rewrite -{1}defB => /imset2P[_ _ /cycleP[i ->] /cycleP[j ->] ->] {uivj}.
+ rewrite !inE sqrB -{-1}[j]odd_double_half.
+ case: (odd j); rewrite -order_dvdn ?ov // ou -def2r -mul2n dvdn_pmul2l //.
+ case/dvdnP=> k ->{i}; apply/orP.
+ rewrite add0n -[j./2]odd_double_half addnC doubleD -!muln2 -mulnA.
+ rewrite -(expg_mod_order v) ov modnMDl; case: (odd _); last first.
+ right; rewrite mulg1 /r -(subnKC n_gt2) expnSr mulnA expgM.
+ by apply: mem_imset; rewrite inE.
+ rewrite (inj_eq (mulIg _)) -expg_mod_order ou -[k]odd_double_half.
+ rewrite addnC -muln2 mulnDl -mulnA def2r modnMDl -ou expg_mod_order.
+ case: (odd k); [left | right]; rewrite ?mul1n ?mul1g //.
+ by apply/imsetP; exists v; rewrite ?inE.
+have nKB: 'N(<<K>>) = B.
+ apply/setP=> b; rewrite !inE -genJ genS // {1}defK conjg_set1 sub1set.
+ have:= Kw; rewrite !inE -!order_dvdn orderJ ow !andbT; apply: contra.
+ case/imsetP=> z _ def_wb; apply/imsetP; exists (z ^ b^-1); rewrite ?inE //.
+ by rewrite -conjXg -def_wb conjgK.
+rewrite -im_quotient card_quotient // nKB -divgS ?subsetT //.
+split; first by rewrite oB defK -orderE ow (mulnA q 2 2) mulnK // mulnC.
+apply: intro_isoGrp => [|rT H].
+ apply/existsP; exists (coset _ u, coset _ v); rewrite /= !xpair_eqE.
+ rewrite -!morphX -?morphJ -?morphV /= ?nKB ?in_setT // uq uv morph1 !eqxx.
+ rewrite -/B -defB -norm_joinEr // quotientY ?nKB ?subsetT //= andbT.
+ rewrite !quotient_cycle /= ?nKB ?in_setT ?eqxx //=.
+ by rewrite -(coset_kerl _ (mem_gen Kw)) -mulgA -expgD v4 mulg1.
+case/existsP=> -[x y] /= /eqP[defH xq y2 xy].
+have ox: #[x] %| #[u] by rewrite ou order_dvdn xq.
+have oy: #[y] %| #[v].
+ by rewrite ov order_dvdn (expgM y 2 2) y2 -expgM mulnC def2r xq.
+have actB: {in <[u]> & <[v]>, morph_act 'J 'J (eltm ox) (eltm oy)}.
+ move=> _ _ /cycleP[i ->] /cycleP[j ->] /=.
+ rewrite conjXg uvj fun_if if_arg fun_if expgVn morphV ?mem_cycle //= !eltmE.
+ rewrite -expgVn -if_arg -fun_if conjXg; congr (_ ^+ i).
+ rewrite -{2}[j]odd_double_half addnC expgD -mul2n expgM y2.
+ rewrite -expgM conjgM (conjgE x) commuteX // mulKg.
+ by case: (odd j); rewrite ?conjg1.
+pose f := sdprodm sdB actB.
+have Kf: 'ker (coset <<K>>) \subset 'ker f.
+ rewrite ker_coset defK cycle_subG /= ker_sdprodm.
+ apply/imset2P; exists (u ^+ r) (v ^+ 2); first exact: mem_cycle.
+ by rewrite inE mem_cycle /= !eltmE y2.
+ by apply: canRL (mulgK _) _; rewrite -mulgA -expgD v4 mulg1.
+have Df: 'dom f \subset 'dom (coset <<K>>) by rewrite /dom nKB subsetT.
+apply/homgP; exists (factm_morphism Kf Df); rewrite morphim_factm /= -/B.
+rewrite -{2}defB morphim_sdprodm // !morphim_cycle ?cycle_id //= !eltm_id.
+by rewrite -norm_joinEr // norms_cycle xy groupV cycle_id.
+Qed.
+
+Lemma card_quaternion : #|'Q_m| = m. Proof. by case defQ. Qed.
+Lemma Grp_quaternion : GrpQ. Proof. by case defQ. Qed.
+
+End Quaternion.
+
+Lemma eq_Mod8_D8 : 'Mod_8 = 'D_8. Proof. by []. Qed.
+
+Section ExtremalStructure.
+
+Variables (gT : finGroupType) (G : {group gT}) (n : nat).
+Implicit Type H : {group gT}.
+
+Let m := (2 ^ n)%N.
+Let q := (2 ^ n.-1)%N.
+Let q_gt0: q > 0. Proof. by rewrite expn_gt0. Qed.
+Let r := (2 ^ n.-2)%N.
+Let r_gt0: r > 0. Proof. by rewrite expn_gt0. Qed.
+
+Let def2qr : n > 1 -> [/\ 2 * q = m, 2 * r = q, q < m & r < q]%N.
+Proof. by rewrite /q /m /r; move/subnKC=> <-; rewrite !ltn_exp2l ?expnS. Qed.
+
+Lemma generators_2dihedral :
+ n > 1 -> G \isog 'D_m ->
+ exists2 xy, extremal_generators G 2 n xy
+ & let: (x, y) := xy in #[y] = 2 /\ x ^ y = x^-1.
+Proof.
+move=> n_gt1; have [def2q _ ltqm _] := def2qr n_gt1.
+case/(isoGrpP _ (Grp_2dihedral n_gt1)); rewrite card_2dihedral // -/ m => oG.
+case/existsP=> -[x y] /=; rewrite -/q => /eqP[defG xq y2 xy].
+have{defG} defG: <[x]> * <[y]> = G.
+ by rewrite -norm_joinEr // norms_cycle xy groupV cycle_id.
+have notXy: y \notin <[x]>.
+ apply: contraL ltqm => Xy; rewrite -leqNgt -oG -defG mulGSid ?cycle_subG //.
+ by rewrite dvdn_leq // order_dvdn xq.
+have oy: #[y] = 2 by exact: nt_prime_order (group1_contra notXy).
+have ox: #[x] = q.
+ apply: double_inj; rewrite -muln2 -oy -mul2n def2q -oG -defG TI_cardMg //.
+ by rewrite setIC prime_TIg ?cycle_subG // -orderE oy.
+exists (x, y) => //=.
+by rewrite oG ox !inE notXy -!cycle_subG /= -defG mulG_subl mulG_subr.
+Qed.
+
+Lemma generators_semidihedral :
+ n > 3 -> G \isog 'SD_m ->
+ exists2 xy, extremal_generators G 2 n xy
+ & let: (x, y) := xy in #[y] = 2 /\ x ^ y = x ^+ r.-1.
+Proof.
+move=> n_gt3; have [def2q _ ltqm _] := def2qr (ltnW (ltnW n_gt3)).
+case/(isoGrpP _ (Grp_semidihedral n_gt3)).
+rewrite card_semidihedral // -/m => oG.
+case/existsP=> -[x y] /=; rewrite -/q -/r => /eqP[defG xq y2 xy].
+have{defG} defG: <[x]> * <[y]> = G.
+ by rewrite -norm_joinEr // norms_cycle xy mem_cycle.
+have notXy: y \notin <[x]>.
+ apply: contraL ltqm => Xy; rewrite -leqNgt -oG -defG mulGSid ?cycle_subG //.
+ by rewrite dvdn_leq // order_dvdn xq.
+have oy: #[y] = 2 by exact: nt_prime_order (group1_contra notXy).
+have ox: #[x] = q.
+ apply: double_inj; rewrite -muln2 -oy -mul2n def2q -oG -defG TI_cardMg //.
+ by rewrite setIC prime_TIg ?cycle_subG // -orderE oy.
+exists (x, y) => //=.
+by rewrite oG ox !inE notXy -!cycle_subG /= -defG mulG_subl mulG_subr.
+Qed.
+
+Lemma generators_quaternion :
+ n > 2 -> G \isog 'Q_m ->
+ exists2 xy, extremal_generators G 2 n xy
+ & let: (x, y) := xy in [/\ #[y] = 4, y ^+ 2 = x ^+ r & x ^ y = x^-1].
+Proof.
+move=> n_gt2; have [def2q def2r ltqm _] := def2qr (ltnW n_gt2).
+case/(isoGrpP _ (Grp_quaternion n_gt2)); rewrite card_quaternion // -/m => oG.
+case/existsP=> -[x y] /=; rewrite -/q -/r => /eqP[defG xq y2 xy].
+have{defG} defG: <[x]> * <[y]> = G.
+ by rewrite -norm_joinEr // norms_cycle xy groupV cycle_id.
+have notXy: y \notin <[x]>.
+ apply: contraL ltqm => Xy; rewrite -leqNgt -oG -defG mulGSid ?cycle_subG //.
+ by rewrite dvdn_leq // order_dvdn xq.
+have ox: #[x] = q.
+ apply/eqP; rewrite eqn_leq dvdn_leq ?order_dvdn ?xq //=.
+ rewrite -(leq_pmul2r (order_gt0 y)) mul_cardG defG oG -def2q mulnAC mulnC.
+ rewrite leq_pmul2r // dvdn_leq ?muln_gt0 ?cardG_gt0 // order_dvdn expgM.
+ by rewrite -order_dvdn order_dvdG //= inE {1}y2 !mem_cycle.
+have oy2: #[y ^+ 2] = 2 by rewrite y2 orderXdiv ox -def2r ?dvdn_mull ?mulnK.
+exists (x, y) => /=; last by rewrite (orderXprime oy2).
+by rewrite oG !inE notXy -!cycle_subG /= -defG mulG_subl mulG_subr.
+Qed.
+
+Variables x y : gT.
+Implicit Type M : {group gT}.
+
+Let X := <[x]>.
+Let Y := <[y]>.
+Let yG := y ^: G.
+Let xyG := (x * y) ^: G.
+Let My := <<yG>>.
+Let Mxy := <<xyG>>.
+
+
+Theorem dihedral2_structure :
+ n > 1 -> extremal_generators G 2 n (x, y) -> G \isog 'D_m ->
+ [/\ [/\ X ><| Y = G, {in G :\: X, forall t, #[t] = 2}
+ & {in X & G :\: X, forall z t, z ^ t = z^-1}],
+ [/\ G ^`(1) = <[x ^+ 2]>, 'Phi(G) = G ^`(1), #|G^`(1)| = r
+ & nil_class G = n.-1],
+ 'Ohm_1(G) = G /\ (forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (2 ^ k)]>),
+ [/\ yG :|: xyG = G :\: X, [disjoint yG & xyG]
+ & forall M, maximal M G = pred3 X My Mxy M]
+ & if n == 2 then (2.-abelem G : Prop) else
+ [/\ 'Z(G) = <[x ^+ r]>, #|'Z(G)| = 2,
+ My \isog 'D_q, Mxy \isog 'D_q
+ & forall U, cyclic U -> U \subset G -> #|G : U| = 2 -> U = X]].
+Proof.
+move=> n_gt1 genG isoG; have [def2q def2r ltqm ltrq] := def2qr n_gt1.
+have [oG Gx ox X'y] := genG; rewrite -/m -/q -/X in oG ox X'y.
+case/extremal_generators_facts: genG; rewrite -/X // => pG maxX nsXG defXY nXY.
+have [sXG nXG]:= andP nsXG; have [Gy notXy]:= setDP X'y.
+have ox2: #[x ^+ 2] = r by rewrite orderXdiv ox -def2r ?dvdn_mulr ?mulKn.
+have oxr: #[x ^+ r] = 2 by rewrite orderXdiv ox -def2r ?dvdn_mull ?mulnK.
+have [[u v] [_ Gu ou U'v] [ov uv]] := generators_2dihedral n_gt1 isoG.
+have defUv: <[u]> :* v = G :\: <[u]>.
+ apply: rcoset_index2; rewrite -?divgS ?cycle_subG //.
+ by rewrite oG -orderE ou -def2q mulnK.
+have invUV: {in <[u]> & <[u]> :* v, forall z t, z ^ t = z^-1}.
+ move=> z t; case/cycleP=> i ->; case/rcosetP=> z'; case/cycleP=> j -> ->{z t}.
+ by rewrite conjgM {2}/conjg commuteX2 // mulKg conjXg uv expgVn.
+have oU': {in <[u]> :* v, forall t, #[t] = 2}.
+ move=> t Uvt; apply: nt_prime_order => //; last first.
+ by case: eqP Uvt => // ->; rewrite defUv !inE group1.
+ case/rcosetP: Uvt => z Uz ->{t}; rewrite expgS {1}(conjgC z) -mulgA.
+ by rewrite invUV ?rcoset_refl // mulKg -(expgS v 1) -ov expg_order.
+have defU: n > 2 -> {in G, forall z, #[z] = q -> <[z]> = <[u]>}.
+ move=> n_gt2 z Gz oz; apply/eqP; rewrite eqEcard -!orderE oz cycle_subG.
+ apply: contraLR n_gt2; rewrite ou leqnn andbT -(ltn_predK n_gt1) => notUz.
+ by rewrite ltnS -(@ltn_exp2l 2) // -/q -oz oU' // defUv inE notUz.
+have n2_abelG: (n > 2) || 2.-abelem G.
+ rewrite ltn_neqAle eq_sym n_gt1; case: eqP => //= n2.
+ apply/abelemP=> //; split=> [|z Gz].
+ by apply: (p2group_abelian pG); rewrite oG pfactorK ?n2.
+ case Uz: (z \in <[u]>); last by rewrite -expg_mod_order oU' // defUv inE Uz.
+ apply/eqP; rewrite -order_dvdn (dvdn_trans (order_dvdG Uz)) // -orderE.
+ by rewrite ou /q n2.
+have{oU'} oX': {in G :\: X, forall t, #[t] = 2}.
+ have [n_gt2 | abelG] := orP n2_abelG; first by rewrite [X]defU // -defUv.
+ move=> t /setDP[Gt notXt]; apply: nt_prime_order (group1_contra notXt) => //.
+ by case/abelemP: abelG => // _ ->.
+have{invUV} invXX': {in X & G :\: X, forall z t, z ^ t = z^-1}.
+ have [n_gt2 | abelG] := orP n2_abelG; first by rewrite [X]defU // -defUv.
+ have [//|cGG oG2] := abelemP _ abelG.
+ move=> t z Xt /setDP[Gz _]; apply/eqP; rewrite eq_sym eq_invg_mul.
+ by rewrite /conjg -(centsP cGG z) // ?mulKg ?[t * t]oG2 ?(subsetP sXG).
+have nXiG k: G \subset 'N(<[x ^+ k]>).
+ apply: char_norm_trans nXG.
+ by rewrite cycle_subgroup_char // cycle_subG mem_cycle.
+have memL i: x ^+ (2 ^ i) \in 'L_i.+1(G).
+ elim: i => // i IHi; rewrite -groupV expnSr expgM invMg.
+ by rewrite -{2}(invXX' _ y) ?mem_cycle ?cycle_id ?mem_commg.
+have defG': G^`(1) = <[x ^+ 2]>.
+ apply/eqP; rewrite eqEsubset cycle_subG (memL 1%N) ?der1_min //=.
+ rewrite (p2group_abelian (quotient_pgroup _ pG)) ?card_quotient //=.
+ rewrite -divgS ?cycle_subG ?groupX // oG -orderE ox2.
+ by rewrite -def2q -def2r mulnA mulnK.
+have defG1: 'Mho^1(G) = <[x ^+ 2]>.
+ apply/eqP; rewrite (MhoE _ pG) eqEsubset !gen_subG sub1set andbC.
+ rewrite mem_gen; last exact: mem_imset.
+ apply/subsetP=> z2; case/imsetP=> z Gz ->{z2}.
+ case Xz: (z \in X); last by rewrite -{1}(oX' z) ?expg_order ?group1 // inE Xz.
+ by case/cycleP: Xz => i ->; rewrite expgAC mem_cycle.
+have defPhi: 'Phi(G) = <[x ^+ 2]>.
+ by rewrite (Phi_joing pG) defG' defG1 (joing_idPl _).
+have def_tG: {in G :\: X, forall t, t ^: G = <[x ^+ 2]> :* t}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ have defJt: {in X, forall z, t ^ z = z ^- 2 * t}.
+ move=> z Xz; rewrite /= invMg -mulgA (conjgC _ t).
+ by rewrite (invXX' _ t) ?groupV ?invgK.
+ have defGt: X * <[t]> = G by rewrite (mulg_normal_maximal nsXG) ?cycle_subG.
+ apply/setP=> tz; apply/imsetP/rcosetP=> [[t'z] | [z]].
+ rewrite -defGt -normC ?cycle_subG ?(subsetP nXG) //.
+ case/imset2P=> _ z /cycleP[j ->] Xz -> -> {tz t'z}.
+ exists (z ^- 2); last by rewrite conjgM {2}/conjg commuteX // mulKg defJt.
+ case/cycleP: Xz => i ->{z}.
+ by rewrite groupV -expgM mulnC expgM mem_cycle.
+ case/cycleP=> i -> -> {z tz}; exists (x ^- i); first by rewrite groupV groupX.
+ by rewrite defJt ?groupV ?mem_cycle // expgVn invgK expgAC.
+have defMt: {in G :\: X, forall t, <[x ^+ 2]> ><| <[t]> = <<t ^: G>>}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ rewrite sdprodEY ?cycle_subG ?(subsetP (nXiG 2)) //; first 1 last.
+ rewrite setIC prime_TIg -?orderE ?oX' // cycle_subG.
+ by apply: contra notXt; apply: subsetP; rewrite cycleX.
+ apply/eqP; have: t \in <<t ^: G>> by rewrite mem_gen ?class_refl.
+ rewrite def_tG // eqEsubset join_subG !cycle_subG !gen_subG => tGt.
+ rewrite tGt -(groupMr _ tGt) mem_gen ?mem_mulg ?cycle_id ?set11 //=.
+ by rewrite mul_subG ?joing_subl // -gen_subG joing_subr.
+have oMt: {in G :\: X, forall t, #|<<t ^: G>>| = q}.
+ move=> t X't /=; rewrite -(sdprod_card (defMt t X't)) -!orderE ox2 oX' //.
+ by rewrite mulnC.
+have sMtG: {in G :\: X, forall t, <<t ^: G>> \subset G}.
+ by move=> t; case/setDP=> Gt _; rewrite gen_subG class_subG.
+have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}.
+ move=> t X't /=; rewrite p_index_maximal -?divgS ?sMtG ?oMt //.
+ by rewrite oG -def2q mulnK.
+have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy.
+have ti_yG_xyG: [disjoint yG & xyG].
+ apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]].
+ rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG.
+ by rewrite -orderE ox2 ox gtnNdvd.
+have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}.
+ by move=> t X't /=; rewrite class_sub_norm // normsD ?normG.
+have defX': yG :|: xyG = G :\: X.
+ apply/eqP; rewrite eqEcard subUset !s_tG_X' //= -(leq_add2l q) -{1}ox orderE.
+ rewrite -/X -{1}(setIidPr sXG) cardsID oG -def2q mul2n -addnn leq_add2l.
+ rewrite -(leq_add2r #|yG :&: xyG|) cardsUI disjoint_setI0 // cards0 addn0.
+ by rewrite /yG /xyG !def_tG // !card_rcoset addnn -mul2n -orderE ox2 def2r.
+split.
+- by rewrite ?sdprodE // setIC // prime_TIg ?cycle_subG // -orderE ?oX'.
+- rewrite defG'; split=> //.
+ apply/eqP; rewrite eqn_leq (leq_trans (nil_class_pgroup pG)); last first.
+ by rewrite oG pfactorK // geq_max leqnn -(subnKC n_gt1).
+ rewrite -(subnKC n_gt1) subn2 ltnNge.
+ rewrite (sameP (lcn_nil_classP _ (pgroup_nil pG)) eqP).
+ by apply/trivgPn; exists (x ^+ r); rewrite ?memL // -order_gt1 oxr.
+- split; last exact: extend_cyclic_Mho.
+ have sX'G1: {subset G :\: X <= 'Ohm_1(G)}.
+ move=> t X't; have [Gt _] := setDP X't.
+ by rewrite (OhmE 1 pG) mem_gen // !inE Gt -(oX' t) //= expg_order.
+ apply/eqP; rewrite eqEsubset Ohm_sub -{1}defXY mulG_subG !cycle_subG.
+ by rewrite -(groupMr _ (sX'G1 y X'y)) !sX'G1.
+- split=> //= H; apply/idP/idP=> [maxH |]; last first.
+ by case/or3P; move/eqP->; rewrite ?maxMt.
+ have [sHG nHG]:= andP (p_maximal_normal pG maxH).
+ have oH: #|H| = q.
+ apply: double_inj; rewrite -muln2 -(p_maximal_index pG maxH) Lagrange //.
+ by rewrite oG -mul2n.
+ rewrite !(eq_sym (gval H)) -eq_sym !eqEcard oH -orderE ox !oMt // !leqnn.
+ case sHX: (H \subset X) => //=; case/subsetPn: sHX => t Ht notXt.
+ have: t \in yG :|: xyG by rewrite defX' inE notXt (subsetP sHG).
+ rewrite !andbT !gen_subG /yG /xyG.
+ by case/setUP; move/class_transr <-; rewrite !class_sub_norm ?Ht ?orbT.
+rewrite eqn_leq n_gt1; case: leqP n2_abelG => //= n_gt2 _.
+have ->: 'Z(G) = <[x ^+ r]>.
+ apply/eqP; rewrite eqEcard andbC -orderE oxr -{1}(setIidPr (center_sub G)).
+ rewrite cardG_gt1 /= meet_center_nil ?(pgroup_nil pG) //; last first.
+ by rewrite -cardG_gt1 oG (leq_trans _ ltqm).
+ apply/subsetP=> t; case/setIP=> Gt cGt.
+ case X't: (t \in G :\: X).
+ move/eqP: (invXX' _ _ (cycle_id x) X't).
+ rewrite /conjg -(centP cGt) // mulKg eq_sym eq_invg_mul -order_eq1 ox2.
+ by rewrite (eqn_exp2l _ 0) // -(subnKC n_gt2).
+ move/idPn: X't; rewrite inE Gt andbT negbK => Xt.
+ have:= Ohm_p_cycle 1 (mem_p_elt pG Gx); rewrite ox pfactorK // subn1 => <-.
+ rewrite (OhmE _ (pgroupS sXG pG)) mem_gen // !inE Xt /=.
+ by rewrite -eq_invg_mul -(invXX' _ y) // /conjg (centP cGt) // mulKg.
+have isoMt: {in G :\: X, forall t, <<t ^: G>> \isog 'D_q}.
+ have n1_gt1: n.-1 > 1 by rewrite -(subnKC n_gt2).
+ move=> t X't /=; rewrite isogEcard card_2dihedral ?oMt // leqnn andbT.
+ rewrite Grp_2dihedral //; apply/existsP; exists (x ^+ 2, t) => /=.
+ have [_ <- nX2T _] := sdprodP (defMt t X't); rewrite norm_joinEr //.
+ rewrite -/q -/r !xpair_eqE eqxx -expgM def2r -ox -{1}(oX' t X't).
+ by rewrite !expg_order !eqxx /= invXX' ?mem_cycle.
+rewrite !isoMt //; split=> // C; case/cyclicP=> z ->{C} sCG iCG.
+rewrite [X]defU // defU -?cycle_subG //.
+by apply: double_inj; rewrite -muln2 -iCG Lagrange // oG -mul2n.
+Qed.
+
+Theorem quaternion_structure :
+ n > 2 -> extremal_generators G 2 n (x, y) -> G \isog 'Q_m ->
+ [/\ [/\ pprod X Y = G, {in G :\: X, forall t, #[t] = 4}
+ & {in X & G :\: X, forall z t, z ^ t = z^-1}],
+ [/\ G ^`(1) = <[x ^+ 2]>, 'Phi(G) = G ^`(1), #|G^`(1)| = r
+ & nil_class G = n.-1],
+ [/\ 'Z(G) = <[x ^+ r]>, #|'Z(G)| = 2,
+ forall u, u \in G -> #[u] = 2 -> u = x ^+ r,
+ 'Ohm_1(G) = <[x ^+ r]> /\ 'Ohm_2(G) = G
+ & forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (2 ^ k)]>],
+ [/\ yG :|: xyG = G :\: X /\ [disjoint yG & xyG]
+ & forall M, maximal M G = pred3 X My Mxy M]
+ & n > 3 ->
+ [/\ My \isog 'Q_q, Mxy \isog 'Q_q
+ & forall U, cyclic U -> U \subset G -> #|G : U| = 2 -> U = X]].
+Proof.
+move=> n_gt2 genG isoG; have [def2q def2r ltqm ltrq] := def2qr (ltnW n_gt2).
+have [oG Gx ox X'y] := genG; rewrite -/m -/q -/X in oG ox X'y.
+case/extremal_generators_facts: genG; rewrite -/X // => pG maxX nsXG defXY nXY.
+have [sXG nXG]:= andP nsXG; have [Gy notXy]:= setDP X'y.
+have oxr: #[x ^+ r] = 2 by rewrite orderXdiv ox -def2r ?dvdn_mull ?mulnK.
+have ox2: #[x ^+ 2] = r by rewrite orderXdiv ox -def2r ?dvdn_mulr ?mulKn.
+have [[u v] [_ Gu ou U'v] [ov v2 uv]] := generators_quaternion n_gt2 isoG.
+have defUv: <[u]> :* v = G :\: <[u]>.
+ apply: rcoset_index2; rewrite -?divgS ?cycle_subG //.
+ by rewrite oG -orderE ou -def2q mulnK.
+have invUV: {in <[u]> & <[u]> :* v, forall z t, z ^ t = z^-1}.
+ move=> z t; case/cycleP=> i ->; case/rcosetP=> ?; case/cycleP=> j -> ->{z t}.
+ by rewrite conjgM {2}/conjg commuteX2 // mulKg conjXg uv expgVn.
+have U'2: {in <[u]> :* v, forall t, t ^+ 2 = u ^+ r}.
+ move=> t; case/rcosetP=> z Uz ->; rewrite expgS {1}(conjgC z) -mulgA.
+ by rewrite invUV ?rcoset_refl // mulKg -(expgS v 1) v2.
+have our: #[u ^+ r] = 2 by rewrite orderXdiv ou -/q -def2r ?dvdn_mull ?mulnK.
+have def_ur: {in G, forall t, #[t] = 2 -> t = u ^+ r}.
+ move=> t Gt /= ot; case Ut: (t \in <[u]>); last first.
+ move/eqP: ot; rewrite eqn_dvd order_dvdn -order_eq1 U'2 ?our //.
+ by rewrite defUv inE Ut.
+ have p2u: 2.-elt u by rewrite /p_elt ou pnat_exp.
+ have: t \in 'Ohm_1(<[u]>).
+ by rewrite (OhmE _ p2u) mem_gen // !inE Ut -order_dvdn ot.
+ rewrite (Ohm_p_cycle _ p2u) ou pfactorK // subn1 -/r cycle_traject our !inE.
+ by rewrite -order_eq1 ot /= mulg1; move/eqP.
+have defU: n > 3 -> {in G, forall z, #[z] = q -> <[z]> = <[u]>}.
+ move=> n_gt3 z Gz oz; apply/eqP; rewrite eqEcard -!orderE oz cycle_subG.
+ rewrite ou leqnn andbT; apply: contraLR n_gt3 => notUz.
+ rewrite -(ltn_predK n_gt2) ltnS -(@ltn_exp2l 2) // -/q -oz.
+ by rewrite (@orderXprime _ 2 2) // U'2 // defUv inE notUz.
+have def_xr: x ^+ r = u ^+ r by apply: def_ur; rewrite ?groupX.
+have X'2: {in G :\: X, forall t, t ^+ 2 = u ^+ r}.
+ case: (ltngtP n 3) => [|n_gt3|n3 t]; first by rewrite ltnNge n_gt2.
+ by rewrite /X defU // -defUv.
+ case/setDP=> Gt notXt.
+ case Ut: (t \in <[u]>); last by rewrite U'2 // defUv inE Ut.
+ rewrite [t ^+ 2]def_ur ?groupX //.
+ have:= order_dvdG Ut; rewrite -orderE ou /q n3 dvdn_divisors ?inE //=.
+ rewrite order_eq1 (negbTE (group1_contra notXt)) /=.
+ case/pred2P=> oz; last by rewrite orderXdiv oz.
+ by rewrite [t]def_ur // -def_xr mem_cycle in notXt.
+have oX': {in G :\: X, forall z, #[z] = 4}.
+ by move=> t X't /=; rewrite (@orderXprime _ 2 2) // X'2.
+have defZ: 'Z(G) = <[x ^+ r]>.
+ apply/eqP; rewrite eqEcard andbC -orderE oxr -{1}(setIidPr (center_sub G)).
+ rewrite cardG_gt1 /= meet_center_nil ?(pgroup_nil pG) //; last first.
+ by rewrite -cardG_gt1 oG (leq_trans _ ltqm).
+ apply/subsetP=> z; case/setIP=> Gz cGz; have [Gv _]:= setDP U'v.
+ case Uvz: (z \in <[u]> :* v).
+ move/eqP: (invUV _ _ (cycle_id u) Uvz).
+ rewrite /conjg -(centP cGz) // mulKg eq_sym eq_invg_mul -(order_dvdn _ 2).
+ by rewrite ou pfactor_dvdn // -(subnKC n_gt2).
+ move/idPn: Uvz; rewrite defUv inE Gz andbT negbK def_xr => Uz.
+ have p_u: 2.-elt u := mem_p_elt pG Gu.
+ suff: z \in 'Ohm_1(<[u]>) by rewrite (Ohm_p_cycle 1 p_u) ou pfactorK // subn1.
+ rewrite (OhmE _ p_u) mem_gen // !inE Uz /= -eq_invg_mul.
+ by rewrite -(invUV _ v) ?rcoset_refl // /conjg (centP cGz) ?mulKg.
+have{invUV} invXX': {in X & G :\: X, forall z t, z ^ t = z^-1}.
+ case: (ltngtP n 3) => [|n_gt3|n3 t z Xt]; first by rewrite ltnNge n_gt2.
+ by rewrite /X defU // -defUv.
+ case/setDP=> Gz notXz; rewrite /q /r n3 /= in oxr ox.
+ suff xz: x ^ z = x^-1 by case/cycleP: Xt => i ->; rewrite conjXg xz expgVn.
+ have: x ^ z \in X by rewrite memJ_norm ?cycle_id ?(subsetP nXG).
+ rewrite invg_expg /X cycle_traject ox !inE /= !mulg1 -order_eq1 orderJ ox /=.
+ case/or3P; move/eqP=> //; last by move/(congr1 order); rewrite orderJ ox oxr.
+ move/conjg_fixP; rewrite (sameP commgP cent1P) cent1C -cent_cycle -/X => cXz.
+ have defXz: X * <[z]> = G by rewrite (mulg_normal_maximal nsXG) ?cycle_subG.
+ have: z \in 'Z(G) by rewrite inE Gz -defXz centM inE cXz cent_cycle cent1id.
+ by rewrite defZ => Xr_z; rewrite (subsetP (cycleX x r)) in notXz.
+have nXiG k: G \subset 'N(<[x ^+ k]>).
+ apply: char_norm_trans nXG.
+ by rewrite cycle_subgroup_char // cycle_subG mem_cycle.
+have memL i: x ^+ (2 ^ i) \in 'L_i.+1(G).
+ elim: i => // i IHi; rewrite -groupV expnSr expgM invMg.
+ by rewrite -{2}(invXX' _ y) ?mem_cycle ?cycle_id ?mem_commg.
+have defG': G^`(1) = <[x ^+ 2]>.
+ apply/eqP; rewrite eqEsubset cycle_subG (memL 1%N) ?der1_min //=.
+ rewrite (p2group_abelian (quotient_pgroup _ pG)) ?card_quotient //=.
+ rewrite -divgS ?cycle_subG ?groupX // oG -orderE ox2.
+ by rewrite -def2q -def2r mulnA mulnK.
+have defG1: 'Mho^1(G) = <[x ^+ 2]>.
+ apply/eqP; rewrite (MhoE _ pG) eqEsubset !gen_subG sub1set andbC.
+ rewrite mem_gen; last exact: mem_imset.
+ apply/subsetP=> z2; case/imsetP=> z Gz ->{z2}.
+ case Xz: (z \in X).
+ by case/cycleP: Xz => i ->; rewrite -expgM mulnC expgM mem_cycle.
+ rewrite (X'2 z) ?inE ?Xz // -def_xr.
+ by rewrite /r -(subnKC n_gt2) expnS expgM mem_cycle.
+have defPhi: 'Phi(G) = <[x ^+ 2]>.
+ by rewrite (Phi_joing pG) defG' defG1 (joing_idPl _).
+have def_tG: {in G :\: X, forall t, t ^: G = <[x ^+ 2]> :* t}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ have defJt: {in X, forall z, t ^ z = z ^- 2 * t}.
+ move=> z Xz; rewrite /= invMg -mulgA (conjgC _ t).
+ by rewrite (invXX' _ t) ?groupV ?invgK.
+ have defGt: X * <[t]> = G by rewrite (mulg_normal_maximal nsXG) ?cycle_subG.
+ apply/setP=> tz; apply/imsetP/rcosetP=> [[t'z] | [z]].
+ rewrite -defGt -normC ?cycle_subG ?(subsetP nXG) //.
+ case/imset2P=> t' z; case/cycleP=> j -> Xz -> -> {tz t'z t'}.
+ exists (z ^- 2); last by rewrite conjgM {2}/conjg commuteX // mulKg defJt.
+ case/cycleP: Xz => i ->{z}.
+ by rewrite groupV -expgM mulnC expgM mem_cycle.
+ case/cycleP=> i -> -> {z tz}; exists (x ^- i); first by rewrite groupV groupX.
+ by rewrite defJt ?groupV ?mem_cycle // expgVn invgK -!expgM mulnC.
+have defMt: {in G :\: X, forall t, <[x ^+ 2]> <*> <[t]> = <<t ^: G>>}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ apply/eqP; have: t \in <<t ^: G>> by rewrite mem_gen ?class_refl.
+ rewrite def_tG // eqEsubset join_subG !cycle_subG !gen_subG => tGt.
+ rewrite tGt -(groupMr _ tGt) mem_gen ?mem_mulg ?cycle_id ?set11 //=.
+ by rewrite mul_subG ?joing_subl // -gen_subG joing_subr.
+have sMtG: {in G :\: X, forall t, <<t ^: G>> \subset G}.
+ by move=> t; case/setDP=> Gt _; rewrite gen_subG class_subG.
+have oMt: {in G :\: X, forall t, #|<<t ^: G>>| = q}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ rewrite -defMt // -(Lagrange (joing_subl _ _)) -orderE ox2 -def2r mulnC.
+ congr (_ * r)%N; rewrite -card_quotient /=; last first.
+ by rewrite defMt // (subset_trans _ (nXiG 2)) ?sMtG.
+ rewrite joingC quotientYidr ?(subset_trans _ (nXiG 2)) ?cycle_subG //.
+ rewrite quotient_cycle ?(subsetP (nXiG 2)) //= -defPhi.
+ rewrite -orderE (abelem_order_p (Phi_quotient_abelem pG)) ?mem_quotient //.
+ apply: contraNneq notXt; move/coset_idr; move/implyP=> /=.
+ by rewrite defPhi ?(subsetP (nXiG 2)) //; apply: subsetP; exact: cycleX.
+have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}.
+ move=> t X't; rewrite /= p_index_maximal -?divgS ?sMtG ?oMt //.
+ by rewrite oG -def2q mulnK.
+have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy.
+have ti_yG_xyG: [disjoint yG & xyG].
+ apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]].
+ rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG.
+ by rewrite -orderE ox2 ox gtnNdvd.
+have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}.
+ by move=> t X't /=; rewrite class_sub_norm // normsD ?normG.
+have defX': yG :|: xyG = G :\: X.
+ apply/eqP; rewrite eqEcard subUset !s_tG_X' //= -(leq_add2l q) -{1}ox orderE.
+ rewrite -/X -{1}(setIidPr sXG) cardsID oG -def2q mul2n -addnn leq_add2l.
+ rewrite -(leq_add2r #|yG :&: xyG|) cardsUI disjoint_setI0 // cards0 addn0.
+ by rewrite /yG /xyG !def_tG // !card_rcoset addnn -mul2n -orderE ox2 def2r.
+rewrite pprodE //; split=> // [|||n_gt3].
+- rewrite defG'; split=> //; apply/eqP; rewrite eqn_leq.
+ rewrite (leq_trans (nil_class_pgroup pG)); last first.
+ by rewrite oG pfactorK // -(subnKC n_gt2).
+ rewrite -(subnKC (ltnW n_gt2)) subn2 ltnNge.
+ rewrite (sameP (lcn_nil_classP _ (pgroup_nil pG)) eqP).
+ by apply/trivgPn; exists (x ^+ r); rewrite ?memL // -order_gt1 oxr.
+- rewrite {2}def_xr defZ; split=> //; last exact: extend_cyclic_Mho.
+ split; apply/eqP; last first.
+ have sX'G2: {subset G :\: X <= 'Ohm_2(G)}.
+ move=> z X'z; have [Gz _] := setDP X'z.
+ by rewrite (OhmE 2 pG) mem_gen // !inE Gz -order_dvdn oX'.
+ rewrite eqEsubset Ohm_sub -{1}defXY mulG_subG !cycle_subG.
+ by rewrite -(groupMr _ (sX'G2 y X'y)) !sX'G2.
+ rewrite eqEsubset (OhmE 1 pG) cycle_subG gen_subG andbC.
+ rewrite mem_gen ?inE ?groupX -?order_dvdn ?oxr //=.
+ apply/subsetP=> t; case/setIP=> Gt; rewrite inE -order_dvdn /=.
+ rewrite dvdn_divisors ?inE //= order_eq1.
+ case/pred2P=> [->|]; first exact: group1.
+ by move/def_ur=> -> //; rewrite def_xr cycle_id.
+- split=> //= H; apply/idP/idP=> [maxH |]; last first.
+ by case/or3P=> /eqP->; rewrite ?maxMt.
+ have [sHG nHG]:= andP (p_maximal_normal pG maxH).
+ have oH: #|H| = q.
+ apply: double_inj; rewrite -muln2 -(p_maximal_index pG maxH) Lagrange //.
+ by rewrite oG -mul2n.
+ rewrite !(eq_sym (gval H)) -eq_sym !eqEcard oH -orderE ox !oMt // !leqnn.
+ case sHX: (H \subset X) => //=; case/subsetPn: sHX => z Hz notXz.
+ have: z \in yG :|: xyG by rewrite defX' inE notXz (subsetP sHG).
+ rewrite !andbT !gen_subG /yG /xyG.
+ by case/setUP=> /class_transr <-; rewrite !class_sub_norm ?Hz ?orbT.
+have isoMt: {in G :\: X, forall z, <<z ^: G>> \isog 'Q_q}.
+ have n1_gt2: n.-1 > 2 by rewrite -(subnKC n_gt3).
+ move=> z X'z /=; rewrite isogEcard card_quaternion ?oMt // leqnn andbT.
+ rewrite Grp_quaternion //; apply/existsP; exists (x ^+ 2, z) => /=.
+ rewrite defMt // -/q -/r !xpair_eqE -!expgM def2r -order_dvdn ox dvdnn.
+ rewrite -expnS prednK; last by rewrite -subn2 subn_gt0.
+ by rewrite X'2 // def_xr !eqxx /= invXX' ?mem_cycle.
+rewrite !isoMt //; split=> // C; case/cyclicP=> z ->{C} sCG iCG.
+rewrite [X]defU // defU -?cycle_subG //.
+by apply: double_inj; rewrite -muln2 -iCG Lagrange // oG -mul2n.
+Qed.
+
+Theorem semidihedral_structure :
+ n > 3 -> extremal_generators G 2 n (x, y) -> G \isog 'SD_m -> #[y] = 2 ->
+ [/\ [/\ X ><| Y = G, #[x * y] = 4
+ & {in X & G :\: X, forall z t, z ^ t = z ^+ r.-1}],
+ [/\ G ^`(1) = <[x ^+ 2]>, 'Phi(G) = G ^`(1), #|G^`(1)| = r
+ & nil_class G = n.-1],
+ [/\ 'Z(G) = <[x ^+ r]>, #|'Z(G)| = 2,
+ 'Ohm_1(G) = My /\ 'Ohm_2(G) = G
+ & forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (2 ^ k)]>],
+ [/\ yG :|: xyG = G :\: X /\ [disjoint yG & xyG]
+ & forall H, maximal H G = pred3 X My Mxy H]
+ & [/\ My \isog 'D_q, Mxy \isog 'Q_q
+ & forall U, cyclic U -> U \subset G -> #|G : U| = 2 -> U = X]].
+Proof.
+move=> n_gt3 genG isoG oy.
+have [def2q def2r ltqm ltrq] := def2qr (ltnW (ltnW n_gt3)).
+have [oG Gx ox X'y] := genG; rewrite -/m -/q -/X in oG ox X'y.
+case/extremal_generators_facts: genG; rewrite -/X // => pG maxX nsXG defXY nXY.
+have [sXG nXG]:= andP nsXG; have [Gy notXy]:= setDP X'y.
+have ox2: #[x ^+ 2] = r by rewrite orderXdiv ox -def2r ?dvdn_mulr ?mulKn.
+have oxr: #[x ^+ r] = 2 by rewrite orderXdiv ox -def2r ?dvdn_mull ?mulnK.
+have [[u v] [_ Gu ou U'v] [ov uv]] := generators_semidihedral n_gt3 isoG.
+have defUv: <[u]> :* v = G :\: <[u]>.
+ apply: rcoset_index2; rewrite -?divgS ?cycle_subG //.
+ by rewrite oG -orderE ou -def2q mulnK.
+have invUV: {in <[u]> & <[u]> :* v, forall z t, z ^ t = z ^+ r.-1}.
+ move=> z t; case/cycleP=> i ->; case/rcosetP=> ?; case/cycleP=> j -> ->{z t}.
+ by rewrite conjgM {2}/conjg commuteX2 // mulKg conjXg uv -!expgM mulnC.
+have [vV yV]: v^-1 = v /\ y^-1 = y by rewrite !invg_expg ov oy.
+have defU: {in G, forall z, #[z] = q -> <[z]> = <[u]>}.
+ move=> z Gz /= oz; apply/eqP; rewrite eqEcard -!orderE oz ou leqnn andbT.
+ apply: contraLR (n_gt3) => notUz; rewrite -leqNgt -(ltn_predK n_gt3) ltnS.
+ rewrite -(@dvdn_Pexp2l 2) // -/q -{}oz order_dvdn expgM (expgS z).
+ have{Gz notUz} [z' Uz' ->{z}]: exists2 z', z' \in <[u]> & z = z' * v.
+ by apply/rcosetP; rewrite defUv inE -cycle_subG notUz Gz.
+ rewrite {2}(conjgC z') invUV ?rcoset_refl // mulgA -{2}vV mulgK -expgS.
+ by rewrite prednK // -expgM mulnC def2r -order_dvdn /q -ou order_dvdG.
+have{invUV} invXX': {in X & G :\: X, forall z t, z ^ t = z ^+ r.-1}.
+ by rewrite /X defU -?defUv.
+have xy2: (x * y) ^+ 2 = x ^+ r.
+ rewrite expgS {2}(conjgC x) invXX' ?cycle_id // mulgA -{2}yV mulgK -expgS.
+ by rewrite prednK.
+have oxy: #[x * y] = 4 by rewrite (@orderXprime _ 2 2) ?xy2.
+have r_gt2: r > 2 by rewrite (ltn_exp2l 1) // -(subnKC n_gt3).
+have coXr1: coprime #[x] (2 ^ (n - 3)).-1.
+ rewrite ox coprime_expl // -(@coprime_pexpl (n - 3)) ?coprimenP ?subn_gt0 //.
+ by rewrite expn_gt0.
+have def2r1: (2 * (2 ^ (n - 3)).-1).+1 = r.-1.
+ rewrite -!subn1 mulnBr -expnS [_.+1]subnSK ?(ltn_exp2l 0) //.
+ by rewrite /r -(subnKC n_gt3).
+have defZ: 'Z(G) = <[x ^+ r]>.
+ apply/eqP; rewrite eqEcard andbC -orderE oxr -{1}(setIidPr (center_sub G)).
+ rewrite cardG_gt1 /= meet_center_nil ?(pgroup_nil pG) //; last first.
+ by rewrite -cardG_gt1 oG (leq_trans _ ltqm).
+ apply/subsetP=> z /setIP[Gz cGz].
+ case X'z: (z \in G :\: X).
+ move/eqP: (invXX' _ _ (cycle_id x) X'z).
+ rewrite /conjg -(centP cGz) // mulKg -def2r1 eq_mulVg1 expgS mulKg mulnC.
+ rewrite -order_dvdn Gauss_dvdr // order_dvdn -order_eq1.
+ by rewrite ox2 -(subnKC r_gt2).
+ move/idPn: X'z; rewrite inE Gz andbT negbK => Xz.
+ have:= Ohm_p_cycle 1 (mem_p_elt pG Gx); rewrite ox pfactorK // subn1 => <-.
+ rewrite (OhmE _ (mem_p_elt pG Gx)) mem_gen // !inE Xz /=.
+ rewrite -(expgK coXr1 Xz) -!expgM mulnCA -order_dvdn dvdn_mull //.
+ rewrite mulnC order_dvdn -(inj_eq (mulgI z)) -expgS mulg1 def2r1.
+ by rewrite -(invXX' z y) // /conjg (centP cGz) ?mulKg.
+have nXiG k: G \subset 'N(<[x ^+ k]>).
+ apply: char_norm_trans nXG.
+ by rewrite cycle_subgroup_char // cycle_subG mem_cycle.
+have memL i: x ^+ (2 ^ i) \in 'L_i.+1(G).
+ elim: i => // i IHi; rewrite -(expgK coXr1 (mem_cycle _ _)) groupX //.
+ rewrite -expgM expnSr -mulnA expgM -(mulKg (x ^+ (2 ^ i)) (_ ^+ _)).
+ by rewrite -expgS def2r1 -(invXX' _ y) ?mem_cycle ?mem_commg.
+have defG': G^`(1) = <[x ^+ 2]>.
+ apply/eqP; rewrite eqEsubset cycle_subG (memL 1%N) ?der1_min //=.
+ rewrite (p2group_abelian (quotient_pgroup _ pG)) ?card_quotient //=.
+ rewrite -divgS ?cycle_subG ?groupX // oG -orderE ox2.
+ by rewrite -def2q -def2r mulnA mulnK.
+have defG1: 'Mho^1(G) = <[x ^+ 2]>.
+ apply/eqP; rewrite (MhoE _ pG) eqEsubset !gen_subG sub1set andbC.
+ rewrite mem_gen; last exact: mem_imset.
+ apply/subsetP=> z2; case/imsetP=> z Gz ->{z2}.
+ case Xz: (z \in X).
+ by case/cycleP: Xz => i ->; rewrite -expgM mulnC expgM mem_cycle.
+ have{Xz Gz} [xi Xxi ->{z}]: exists2 xi, xi \in X & z = xi * y.
+ have Uvy: y \in <[u]> :* v by rewrite defUv -(defU x).
+ apply/rcosetP; rewrite /X defU // (rcoset_transl Uvy) defUv.
+ by rewrite inE -(defU x) ?Xz.
+ rewrite expn1 expgS {2}(conjgC xi) -{2}[y]/(y ^+ 2.-1) -{1}oy -invg_expg.
+ rewrite mulgA mulgK invXX' // -expgS prednK // /r -(subnKC n_gt3) expnS.
+ by case/cycleP: Xxi => i ->; rewrite -expgM mulnCA expgM mem_cycle.
+have defPhi: 'Phi(G) = <[x ^+ 2]>.
+ by rewrite (Phi_joing pG) defG' defG1 (joing_idPl _).
+have def_tG: {in G :\: X, forall t, t ^: G = <[x ^+ 2]> :* t}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ have defJt: {in X, forall z, t ^ z = z ^+ r.-2 * t}.
+ move=> z Xz /=; rewrite -(mulKg z (z ^+ _)) -expgS -subn2.
+ have X'tV: t^-1 \in G :\: X by rewrite inE !groupV notXt.
+ by rewrite subnSK 1?ltnW // subn1 -(invXX' _ t^-1) // -mulgA -conjgCV.
+ have defGt: X * <[t]> = G by rewrite (mulg_normal_maximal nsXG) ?cycle_subG.
+ apply/setP=> tz; apply/imsetP/rcosetP=> [[t'z] | [z]].
+ rewrite -defGt -normC ?cycle_subG ?(subsetP nXG) //.
+ case/imset2P=> t' z; case/cycleP=> j -> Xz -> -> {t' t'z tz}.
+ exists (z ^+ r.-2); last first.
+ by rewrite conjgM {2}/conjg commuteX // mulKg defJt.
+ case/cycleP: Xz => i ->{z}.
+ by rewrite -def2r1 -expgM mulnCA expgM mem_cycle.
+ case/cycleP=> i -> -> {z tz}.
+ exists (x ^+ (i * expg_invn X (2 ^ (n - 3)).-1)); first by rewrite groupX.
+ rewrite defJt ?mem_cycle // -def2r1 -!expgM.
+ by rewrite mulnAC mulnA mulnC muln2 !expgM expgK ?mem_cycle.
+have defMt: {in G :\: X, forall t, <[x ^+ 2]> <*> <[t]> = <<t ^: G>>}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ apply/eqP; have: t \in <<t ^: G>> by rewrite mem_gen ?class_refl.
+ rewrite def_tG // eqEsubset join_subG !cycle_subG !gen_subG => tGt.
+ rewrite tGt -(groupMr _ tGt) mem_gen ?mem_mulg ?cycle_id ?set11 //=.
+ by rewrite mul_subG ?joing_subl // -gen_subG joing_subr.
+have sMtG: {in G :\: X, forall t, <<t ^: G>> \subset G}.
+ by move=> t; case/setDP=> Gt _; rewrite gen_subG class_subG.
+have oMt: {in G :\: X, forall t, #|<<t ^: G>>| = q}.
+ move=> t X't; have [Gt notXt] := setDP X't.
+ rewrite -defMt // -(Lagrange (joing_subl _ _)) -orderE ox2 -def2r mulnC.
+ congr (_ * r)%N; rewrite -card_quotient /=; last first.
+ by rewrite defMt // (subset_trans _ (nXiG 2)) ?sMtG.
+ rewrite joingC quotientYidr ?(subset_trans _ (nXiG 2)) ?cycle_subG //.
+ rewrite quotient_cycle ?(subsetP (nXiG 2)) //= -defPhi -orderE.
+ rewrite (abelem_order_p (Phi_quotient_abelem pG)) ?mem_quotient //.
+ apply: contraNneq notXt; move/coset_idr; move/implyP=> /=.
+ by rewrite /= defPhi (subsetP (nXiG 2)) //; apply: subsetP; exact: cycleX.
+have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}.
+ move=> t X't /=; rewrite p_index_maximal -?divgS ?sMtG ?oMt //.
+ by rewrite oG -def2q mulnK.
+have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy.
+have ti_yG_xyG: [disjoint yG & xyG].
+ apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]].
+ rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG.
+ by rewrite -orderE ox2 ox gtnNdvd.
+have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}.
+ by move=> t X't /=; rewrite class_sub_norm // normsD ?normG.
+have defX': yG :|: xyG = G :\: X.
+ apply/eqP; rewrite eqEcard subUset !s_tG_X' //= -(leq_add2l q) -{1}ox orderE.
+ rewrite -/X -{1}(setIidPr sXG) cardsID oG -def2q mul2n -addnn leq_add2l.
+ rewrite -(leq_add2r #|yG :&: xyG|) cardsUI disjoint_setI0 // cards0 addn0.
+ by rewrite /yG /xyG !def_tG // !card_rcoset addnn -mul2n -orderE ox2 def2r.
+split.
+- by rewrite sdprodE // setIC prime_TIg ?cycle_subG // -orderE oy.
+- rewrite defG'; split=> //.
+ apply/eqP; rewrite eqn_leq (leq_trans (nil_class_pgroup pG)); last first.
+ by rewrite oG pfactorK // -(subnKC n_gt3).
+ rewrite -(subnKC (ltnW (ltnW n_gt3))) subn2 ltnNge.
+ rewrite (sameP (lcn_nil_classP _ (pgroup_nil pG)) eqP).
+ by apply/trivgPn; exists (x ^+ r); rewrite ?memL // -order_gt1 oxr.
+- rewrite defZ; split=> //; last exact: extend_cyclic_Mho.
+ split; apply/eqP; last first.
+ have sX'G2: {subset G :\: X <= 'Ohm_2(G)}.
+ move=> t X't; have [Gt _] := setDP X't; rewrite -defX' in X't.
+ rewrite (OhmE 2 pG) mem_gen // !inE Gt -order_dvdn.
+ by case/setUP: X't; case/imsetP=> z _ ->; rewrite orderJ ?oy ?oxy.
+ rewrite eqEsubset Ohm_sub -{1}defXY mulG_subG !cycle_subG.
+ by rewrite -(groupMr _ (sX'G2 y X'y)) !sX'G2.
+ rewrite eqEsubset andbC gen_subG class_sub_norm ?gFnorm //.
+ rewrite (OhmE 1 pG) mem_gen ?inE ?Gy -?order_dvdn ?oy // gen_subG /= -/My.
+ apply/subsetP=> t; rewrite !inE; case/andP=> Gt t2.
+ have pX := pgroupS sXG pG.
+ case Xt: (t \in X).
+ have: t \in 'Ohm_1(X) by rewrite (OhmE 1 pX) mem_gen // !inE Xt.
+ apply: subsetP; rewrite (Ohm_p_cycle 1 pX) ox pfactorK //.
+ rewrite -(subnKC n_gt3) expgM (subset_trans (cycleX _ _)) //.
+ by rewrite /My -defMt ?joing_subl.
+ have{Xt}: t \in yG :|: xyG by rewrite defX' inE Xt.
+ case/setUP; first exact: mem_gen.
+ by case/imsetP=> z _ def_t; rewrite -order_dvdn def_t orderJ oxy in t2.
+- split=> //= H; apply/idP/idP=> [maxH |]; last first.
+ by case/or3P=> /eqP->; rewrite ?maxMt.
+ have [sHG nHG]:= andP (p_maximal_normal pG maxH).
+ have oH: #|H| = q.
+ apply: double_inj; rewrite -muln2 -(p_maximal_index pG maxH) Lagrange //.
+ by rewrite oG -mul2n.
+ rewrite !(eq_sym (gval H)) -eq_sym !eqEcard oH -orderE ox !oMt // !leqnn.
+ case sHX: (H \subset X) => //=; case/subsetPn: sHX => t Ht notXt.
+ have: t \in yG :|: xyG by rewrite defX' inE notXt (subsetP sHG).
+ rewrite !andbT !gen_subG /yG /xyG.
+ by case/setUP=> /class_transr <-; rewrite !class_sub_norm ?Ht ?orbT.
+have n1_gt2: n.-1 > 2 by [rewrite -(subnKC n_gt3)]; have n1_gt1 := ltnW n1_gt2.
+rewrite !isogEcard card_2dihedral ?card_quaternion ?oMt // leqnn !andbT.
+have invX2X': {in G :\: X, forall t, x ^+ 2 ^ t == x ^- 2}.
+ move=> t X't; rewrite /= invXX' ?mem_cycle // eq_sym eq_invg_mul -expgS.
+ by rewrite prednK // -order_dvdn ox2.
+ rewrite Grp_2dihedral ?Grp_quaternion //; split=> [||C].
+- apply/existsP; exists (x ^+ 2, y); rewrite /= defMt // !xpair_eqE.
+ by rewrite -!expgM def2r -!order_dvdn ox oy dvdnn eqxx /= invX2X'.
+- apply/existsP; exists (x ^+ 2, x * y); rewrite /= defMt // !xpair_eqE.
+ rewrite -!expgM def2r -order_dvdn ox xy2 dvdnn eqxx invX2X' //=.
+ by rewrite andbT /r -(subnKC n_gt3).
+case/cyclicP=> z ->{C} sCG iCG; rewrite [X]defU // defU -?cycle_subG //.
+by apply: double_inj; rewrite -muln2 -iCG Lagrange // oG -mul2n.
+Qed.
+
+End ExtremalStructure.
+
+Section ExtremalClass.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Inductive extremal_group_type :=
+ ModularGroup | Dihedral | SemiDihedral | Quaternion | NotExtremal.
+
+Definition index_extremal_group_type c :=
+ match c with
+ | ModularGroup => 0
+ | Dihedral => 1
+ | SemiDihedral => 2
+ | Quaternion => 3
+ | NotExtremal => 4
+ end%N.
+
+Definition enum_extremal_groups :=
+ [:: ModularGroup; Dihedral; SemiDihedral; Quaternion].
+
+Lemma cancel_index_extremal_groups :
+ cancel index_extremal_group_type (nth NotExtremal enum_extremal_groups).
+Proof. by case. Qed.
+Local Notation extgK := cancel_index_extremal_groups.
+
+Import choice.
+
+Definition extremal_group_eqMixin := CanEqMixin extgK.
+Canonical extremal_group_eqType := EqType _ extremal_group_eqMixin.
+Definition extremal_group_choiceMixin := CanChoiceMixin extgK.
+Canonical extremal_group_choiceType := ChoiceType _ extremal_group_choiceMixin.
+Definition extremal_group_countMixin := CanCountMixin extgK.
+Canonical extremal_group_countType := CountType _ extremal_group_countMixin.
+Lemma bound_extremal_groups (c : extremal_group_type) : pickle c < 6.
+Proof. by case: c. Qed.
+Definition extremal_group_finMixin := Finite.CountMixin bound_extremal_groups.
+Canonical extremal_group_finType := FinType _ extremal_group_finMixin.
+
+Definition extremal_class (A : {set gT}) :=
+ let m := #|A| in let p := pdiv m in let n := logn p m in
+ if (n > 1) && (A \isog 'D_(2 ^ n)) then Dihedral else
+ if (n > 2) && (A \isog 'Q_(2 ^ n)) then Quaternion else
+ if (n > 3) && (A \isog 'SD_(2 ^ n)) then SemiDihedral else
+ if (n > 2) && (A \isog 'Mod_(p ^ n)) then ModularGroup else
+ NotExtremal.
+
+Definition extremal2 A := extremal_class A \in behead enum_extremal_groups.
+
+Lemma dihedral_classP :
+ extremal_class G = Dihedral <-> (exists2 n, n > 1 & G \isog 'D_(2 ^ n)).
+Proof.
+rewrite /extremal_class; split=> [ | [n n_gt1 isoG]].
+ by move: (logn _ _) => n; do 4?case: ifP => //; case/andP; exists n.
+rewrite (card_isog isoG) card_2dihedral // -(ltn_predK n_gt1) pdiv_pfactor //.
+by rewrite pfactorK // (ltn_predK n_gt1) n_gt1 isoG.
+Qed.
+
+Lemma quaternion_classP :
+ extremal_class G = Quaternion <-> (exists2 n, n > 2 & G \isog 'Q_(2 ^ n)).
+Proof.
+rewrite /extremal_class; split=> [ | [n n_gt2 isoG]].
+ by move: (logn _ _) => n; do 4?case: ifP => //; case/andP; exists n.
+rewrite (card_isog isoG) card_quaternion // -(ltn_predK n_gt2) pdiv_pfactor //.
+rewrite pfactorK // (ltn_predK n_gt2) n_gt2 isoG.
+case: andP => // [[n_gt1 isoGD]].
+have [[x y] genG [oy _ _]]:= generators_quaternion n_gt2 isoG.
+have [_ _ _ X'y] := genG.
+by case/dihedral2_structure: genG oy => // [[_ ->]].
+Qed.
+
+Lemma semidihedral_classP :
+ extremal_class G = SemiDihedral <-> (exists2 n, n > 3 & G \isog 'SD_(2 ^ n)).
+Proof.
+rewrite /extremal_class; split=> [ | [n n_gt3 isoG]].
+ by move: (logn _ _) => n; do 4?case: ifP => //; case/andP; exists n.
+rewrite (card_isog isoG) card_semidihedral //.
+rewrite -(ltn_predK n_gt3) pdiv_pfactor // pfactorK // (ltn_predK n_gt3) n_gt3.
+have [[x y] genG [oy _]]:= generators_semidihedral n_gt3 isoG.
+have [_ Gx _ X'y]:= genG.
+case: andP => [[n_gt1 isoGD]|_].
+ have [[_ oxy _ _] _ _ _]:= semidihedral_structure n_gt3 genG isoG oy.
+ case: (dihedral2_structure n_gt1 genG isoGD) oxy => [[_ ->]] //.
+ by rewrite !inE !groupMl ?cycle_id in X'y *.
+case: andP => // [[n_gt2 isoGQ]|]; last by rewrite isoG.
+by case: (quaternion_structure n_gt2 genG isoGQ) oy => [[_ ->]].
+Qed.
+
+Lemma odd_not_extremal2 : odd #|G| -> ~~ extremal2 G.
+Proof.
+rewrite /extremal2 /extremal_class; case: logn => // n'.
+case: andP => [[n_gt1 isoG] | _].
+ by rewrite (card_isog isoG) card_2dihedral ?odd_exp.
+case: andP => [[n_gt2 isoG] | _].
+ by rewrite (card_isog isoG) card_quaternion ?odd_exp.
+case: andP => [[n_gt3 isoG] | _].
+ by rewrite (card_isog isoG) card_semidihedral ?odd_exp.
+by case: ifP.
+Qed.
+
+Lemma modular_group_classP :
+ extremal_class G = ModularGroup
+ <-> (exists2 p, prime p &
+ exists2 n, n >= (p == 2) + 3 & G \isog 'Mod_(p ^ n)).
+Proof.
+rewrite /extremal_class; split=> [ | [p p_pr [n n_gt23 isoG]]].
+ move: (pdiv _) => p; set n := logn p _; do 4?case: ifP => //.
+ case/andP=> n_gt2 isoG _ _; rewrite ltnW //= => not_isoG _.
+ exists p; first by move: n_gt2; rewrite /n lognE; case (prime p).
+ exists n => //; case: eqP => // p2; rewrite ltn_neqAle; case: eqP => // n3.
+ by case/idP: not_isoG; rewrite p2 -n3 in isoG *.
+have n_gt2 := leq_trans (leq_addl _ _) n_gt23; have n_gt1 := ltnW n_gt2.
+have n_gt0 := ltnW n_gt1; have def_n := prednK n_gt0.
+have [[x y] genG mod_xy] := generators_modular_group p_pr n_gt2 isoG.
+case/modular_group_structure: (genG) => // _ _ [_ _ nil2G] _ _.
+have [oG _ _ _] := genG; have [oy _] := mod_xy.
+rewrite oG -def_n pdiv_pfactor // def_n pfactorK // n_gt1 n_gt2 {}isoG /=.
+case: (ltngtP p 2) => [|p_gt2|p2]; first by rewrite ltnNge prime_gt1.
+ rewrite !(isog_sym G) !isogEcard card_2dihedral ?card_quaternion //= oG.
+ rewrite leq_exp2r // leqNgt p_gt2 !andbF; case: and3P=> // [[n_gt3 _]].
+ by rewrite card_semidihedral // leq_exp2r // leqNgt p_gt2.
+rewrite p2 in genG oy n_gt23; rewrite n_gt23.
+have: nil_class G <> n.-1.
+ by apply/eqP; rewrite neq_ltn -ltnS nil2G def_n n_gt23.
+case: ifP => [isoG | _]; first by case/dihedral2_structure: genG => // _ [].
+case: ifP => [isoG | _]; first by case/quaternion_structure: genG => // _ [].
+by case: ifP => // isoG; case/semidihedral_structure: genG => // _ [].
+Qed.
+
+End ExtremalClass.
+
+Theorem extremal2_structure (gT : finGroupType) (G : {group gT}) n x y :
+ let cG := extremal_class G in
+ let m := (2 ^ n)%N in let q := (2 ^ n.-1)%N in let r := (2 ^ n.-2)%N in
+ let X := <[x]> in let yG := y ^: G in let xyG := (x * y) ^: G in
+ let My := <<yG>> in let Mxy := <<xyG>> in
+ extremal_generators G 2 n (x, y) ->
+ extremal2 G -> (cG == SemiDihedral) ==> (#[y] == 2) ->
+ [/\ [/\ (if cG == Quaternion then pprod X <[y]> else X ><| <[y]>) = G,
+ if cG == SemiDihedral then #[x * y] = 4 else
+ {in G :\: X, forall z, #[z] = (if cG == Dihedral then 2 else 4)},
+ if cG != Quaternion then True else
+ {in G, forall z, #[z] = 2 -> z = x ^+ r}
+ & {in X & G :\: X, forall t z,
+ t ^ z = (if cG == SemiDihedral then t ^+ r.-1 else t^-1)}],
+ [/\ G ^`(1) = <[x ^+ 2]>, 'Phi(G) = G ^`(1), #|G^`(1)| = r
+ & nil_class G = n.-1],
+ [/\ if n > 2 then 'Z(G) = <[x ^+ r]> /\ #|'Z(G)| = 2 else 2.-abelem G,
+ 'Ohm_1(G) = (if cG == Quaternion then <[x ^+ r]> else
+ if cG == SemiDihedral then My else G),
+ 'Ohm_2(G) = G
+ & forall k, k > 0 -> 'Mho^k(G) = <[x ^+ (2 ^ k)]>],
+ [/\ yG :|: xyG = G :\: X, [disjoint yG & xyG]
+ & forall H : {group gT}, maximal H G = (gval H \in pred3 X My Mxy)]
+ & if n <= (cG == Quaternion) + 2 then True else
+ [/\ forall U, cyclic U -> U \subset G -> #|G : U| = 2 -> U = X,
+ if cG == Quaternion then My \isog 'Q_q else My \isog 'D_q,
+ extremal_class My = (if cG == Quaternion then cG else Dihedral),
+ if cG == Dihedral then Mxy \isog 'D_q else Mxy \isog 'Q_q
+ & extremal_class Mxy = (if cG == Dihedral then cG else Quaternion)]].
+Proof.
+move=> cG m q r X yG xyG My Mxy genG; have [oG _ _ _] := genG.
+have logG: logn (pdiv #|G|) #|G| = n by rewrite oG pfactorKpdiv.
+rewrite /extremal2 -/cG; do [rewrite {1}/extremal_class /= {}logG] in cG *.
+case: ifP => [isoG | _] in cG * => [_ _ /=|].
+ case/andP: isoG => n_gt1 isoG.
+ have:= dihedral2_structure n_gt1 genG isoG; rewrite -/X -/q -/r -/yG -/xyG.
+ case=> [[defG oX' invXX'] nilG [defOhm defMho] maxG defZ].
+ rewrite eqn_leq n_gt1 andbT add0n in defZ *; split=> //.
+ split=> //; first by case: leqP defZ => // _ [].
+ by apply/eqP; rewrite eqEsubset Ohm_sub -{1}defOhm Ohm_leq.
+ case: leqP defZ => // n_gt2 [_ _ isoMy isoMxy defX].
+ have n1_gt1: n.-1 > 1 by rewrite -(subnKC n_gt2).
+ by split=> //; apply/dihedral_classP; exists n.-1.
+case: ifP => [isoG | _] in cG * => [_ _ /=|].
+ case/andP: isoG => n_gt2 isoG; rewrite n_gt2 add1n.
+ have:= quaternion_structure n_gt2 genG isoG; rewrite -/X -/q -/r -/yG -/xyG.
+ case=> [[defG oX' invXX'] nilG [defZ oZ def2 [-> ->] defMho]].
+ case=> [[-> ->] maxG] isoM; split=> //.
+ case: leqP isoM => // n_gt3 [//|isoMy isoMxy defX].
+ have n1_gt2: n.-1 > 2 by rewrite -(subnKC n_gt3).
+ by split=> //; apply/quaternion_classP; exists n.-1.
+do [case: ifP => [isoG | _]; last by case: ifP] in cG * => /= _; move/eqnP=> oy.
+case/andP: isoG => n_gt3 isoG; rewrite (leqNgt n) (ltnW n_gt3) /=.
+have n1_gt2: n.-1 > 2 by rewrite -(subnKC n_gt3).
+have:= semidihedral_structure n_gt3 genG isoG oy.
+rewrite -/X -/q -/r -/yG -/xyG -/My -/Mxy.
+case=> [[defG oxy invXX'] nilG [defZ oZ [-> ->] defMho] [[defX' tiX'] maxG]].
+case=> isoMy isoMxy defX; do 2!split=> //.
+ by apply/dihedral_classP; exists n.-1; first exact: ltnW.
+by apply/quaternion_classP; exists n.-1.
+Qed.
+
+(* This is Aschbacher (23.4). *)
+Lemma maximal_cycle_extremal gT p (G X : {group gT}) :
+ p.-group G -> ~~ abelian G -> cyclic X -> X \subset G -> #|G : X| = p ->
+ (extremal_class G == ModularGroup) || (p == 2) && extremal2 G.
+Proof.
+move=> pG not_cGG cycX sXG iXG; rewrite /extremal2; set cG := extremal_class G.
+have [|p_pr _ _] := pgroup_pdiv pG.
+ by case: eqP not_cGG => // ->; rewrite abelian1.
+have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1.
+have [n oG] := p_natP pG; have n_gt2: n > 2.
+ apply: contraR not_cGG; rewrite -leqNgt => n_le2.
+ by rewrite (p2group_abelian pG) // oG pfactorK.
+have def_n := subnKC n_gt2; have n_gt1 := ltnW n_gt2; have n_gt0 := ltnW n_gt1.
+pose q := (p ^ n.-1)%N; pose r := (p ^ n.-2)%N.
+have q_gt1: q > 1 by rewrite (ltn_exp2l 0) // -(subnKC n_gt2).
+have r_gt0: r > 0 by rewrite expn_gt0 p_gt0.
+have def_pr: (p * r)%N = q by rewrite /q /r -def_n.
+have oX: #|X| = q by rewrite -(divg_indexS sXG) oG iXG /q -def_n mulKn.
+have ntX: X :!=: 1 by rewrite -cardG_gt1 oX.
+have maxX: maximal X G by rewrite p_index_maximal ?iXG.
+have nsXG: X <| G := p_maximal_normal pG maxX; have [_ nXG] := andP nsXG.
+have cXX: abelian X := cyclic_abelian cycX.
+have scXG: 'C_G(X) = X.
+ apply/eqP; rewrite eqEsubset subsetI sXG -abelianE cXX !andbT.
+ apply: contraR not_cGG; case/subsetPn=> y; case/setIP=> Gy cXy notXy.
+ rewrite -!cycle_subG in Gy notXy; rewrite -(mulg_normal_maximal nsXG _ Gy) //.
+ by rewrite abelianM cycle_abelian cyclic_abelian ?cycle_subG.
+have [x defX] := cyclicP cycX; have pX := pgroupS sXG pG.
+have Xx: x \in X by [rewrite defX cycle_id]; have Gx := subsetP sXG x Xx.
+have [ox p_x]: #[x] = q /\ p.-elt x by rewrite defX in pX oX.
+pose Z := <[x ^+ r]>.
+have defZ: Z = 'Ohm_1(X) by rewrite defX (Ohm_p_cycle _ p_x) ox subn1 pfactorK.
+have oZ: #|Z| = p by rewrite -orderE orderXdiv ox -def_pr ?dvdn_mull ?mulnK.
+have cGZ: Z \subset 'C(G).
+ have nsZG: Z <| G by rewrite defZ (char_normal_trans (Ohm_char 1 _)).
+ move/implyP: (meet_center_nil (pgroup_nil pG) nsZG).
+ rewrite -cardG_gt1 oZ p_gt1 setIA (setIidPl (normal_sub nsZG)).
+ by apply: contraR; move/prime_TIg=> -> //; rewrite oZ.
+have X_Gp y: y \in G -> y ^+ p \in X.
+ move=> Gy; have nXy: y \in 'N(X) := subsetP nXG y Gy.
+ rewrite coset_idr ?groupX // morphX //; apply/eqP.
+ by rewrite -order_dvdn -iXG -card_quotient // order_dvdG ?mem_quotient.
+have [y X'y]: exists2 y, y \in G :\: X &
+ (p == 2) + 3 <= n /\ x ^ y = x ^+ r.+1 \/ p = 2 /\ x * x ^ y \in Z.
+- have [y Gy notXy]: exists2 y, y \in G & y \notin X.
+ by apply/subsetPn; rewrite proper_subn ?(maxgroupp maxX).
+ have nXy: y \in 'N(X) := subsetP nXG y Gy; pose ay := conj_aut X y.
+ have oay: #[ay] = p.
+ apply: nt_prime_order => //.
+ by rewrite -morphX // mker // ker_conj_aut (subsetP cXX) ?X_Gp.
+ rewrite (sameP eqP (kerP _ nXy)) ker_conj_aut.
+ by apply: contra notXy => cXy; rewrite -scXG inE Gy.
+ have [m []]:= cyclic_pgroup_Aut_structure pX cycX ntX.
+ set Ap := 'O_p(_); case=> def_m [m1 _] [m_inj _] _ _ _.
+ have sylAp: p.-Sylow(Aut X) Ap.
+ by rewrite nilpotent_pcore_Hall // abelian_nil // Aut_cyclic_abelian.
+ have Ap1ay: ay \in 'Ohm_1(Ap).
+ rewrite (OhmE _ (pcore_pgroup _ _)) mem_gen // !inE -order_dvdn oay dvdnn.
+ rewrite (mem_normal_Hall sylAp) ?pcore_normal ?Aut_aut //.
+ by rewrite /p_elt oay pnat_id.
+ rewrite {1}oX pfactorK // -{1}def_n /=.
+ have [p2 | odd_p] := even_prime p_pr; last first.
+ rewrite (sameP eqP (prime_oddPn p_pr)) odd_p n_gt2.
+ case=> _ [_ _ _] [_ _ [s [As os m_s defAp1]]].
+ have [j def_s]: exists j, s = ay ^+ j.
+ apply/cycleP; rewrite -cycle_subG subEproper eq_sym eqEcard -!orderE.
+ by rewrite -defAp1 cycle_subG Ap1ay oay os leqnn .
+ exists (y ^+ j); last first.
+ left; rewrite -(norm_conj_autE _ Xx) ?groupX // morphX // -def_s.
+ by rewrite -def_m // m_s expg_znat // oX pfactorK ?eqxx.
+ rewrite -scXG !inE groupX //= andbT -ker_conj_aut !inE morphX // -def_s.
+ rewrite andbC -(inj_in_eq m_inj) ?group1 // m_s m1 oX pfactorK // -/r.
+ rewrite mulrSr -subr_eq0 addrK -val_eqE /= val_Zp_nat //.
+ by rewrite [_ == 0%N]dvdn_Pexp2l // -def_n ltnn.
+ rewrite {1}p2 /= => [[t [At ot m_t]]]; rewrite {1}oX pfactorK // -{1}def_n.
+ rewrite eqSS subn_eq0 => defA; exists y; rewrite ?inE ?notXy //.
+ rewrite p2 -(norm_conj_autE _ Xx) //= -/ay -def_m ?Aut_aut //.
+ case Tay: (ay \in <[t]>).
+ rewrite cycle2g // !inE -order_eq1 oay p2 /= in Tay.
+ by right; rewrite (eqP Tay) m_t expg_zneg // mulgV group1.
+ case: leqP defA => [_ defA|le3n [a [Aa _ _ defA [s [As os m_s m_st defA1]]]]].
+ by rewrite -defA Aut_aut in Tay.
+ have: ay \in [set s; s * t].
+ have: ay \in 'Ohm_1(Aut X) := subsetP (OhmS 1 (pcore_sub _ _)) ay Ap1ay.
+ case/dprodP: (Ohm_dprod 1 defA) => _ <- _ _.
+ rewrite defA1 (@Ohm_p_cycle _ _ 2) /p_elt ot //= expg1 cycle2g //.
+ by rewrite mulUg mul1g inE Tay cycle2g // mulgU mulg1 mulg_set1.
+ case/set2P=> ->; [left | right].
+ by rewrite ?le3n m_s expg_znat // oX pfactorK // -p2.
+ by rewrite m_st expg_znat // oX pfactorK // -p2 -/r -expgS prednK ?cycle_id.
+have [Gy notXy] := setDP X'y; have nXy := subsetP nXG y Gy.
+have defG j: <[x]> <*> <[x ^+ j * y]> = G.
+ rewrite -defX -genM_join.
+ by rewrite (mulg_normal_maximal nsXG) ?cycle_subG ?groupMl ?groupX ?genGid.
+have[i def_yp]: exists i, y ^- p = x ^+ i.
+ by apply/cycleP; rewrite -defX groupV X_Gp.
+have p_i: p %| i.
+ apply: contraR notXy; rewrite -prime_coprime // => co_p_j.
+ have genX: generator X (y ^- p).
+ by rewrite def_yp defX generator_coprime ox coprime_expl.
+ rewrite -scXG (setIidPl _) // centsC ((X :=P: _) genX) cycle_subG groupV.
+ rewrite /= -(defG 0%N) mul1g centY inE -defX (subsetP cXX) ?X_Gp //.
+ by rewrite (subsetP (cycle_abelian y)) ?mem_cycle.
+case=> [[n_gt23 xy] | [p2 Z_xxy]].
+ suffices ->: cG = ModularGroup by []; apply/modular_group_classP.
+ exists p => //; exists n => //; rewrite isogEcard card_modular_group //.
+ rewrite oG leqnn andbT Grp_modular_group // -/q -/r.
+ have{i def_yp p_i} [i def_yp]: exists i, y ^- p = x ^+ i ^+ p.
+ by case/dvdnP: p_i => j def_i; exists j; rewrite -expgM -def_i.
+ have Zyx: [~ y, x] \in Z.
+ by rewrite -groupV invg_comm commgEl xy expgS mulKg cycle_id.
+ have def_yxj j: [~ y, x ^+ j] = [~ y, x] ^+ j.
+ by rewrite commgX /commute ?(centsP cGZ _ Zyx).
+ have Zyxj j: [~ y, x ^+ j] \in Z by rewrite def_yxj groupX.
+ have x_xjy j: x ^ (x ^+ j * y) = x ^+ r.+1.
+ by rewrite conjgM {2}/conjg commuteX //= mulKg.
+ have [cyxi | not_cyxi] := eqVneq ([~ y, x ^+ i] ^+ 'C(p, 2)) 1.
+ apply/existsP; exists (x, x ^+ i * y); rewrite /= !xpair_eqE.
+ rewrite defG x_xjy -order_dvdn ox dvdnn !eqxx andbT /=.
+ rewrite expMg_Rmul /commute ?(centsP cGZ _ (Zyxj _)) ?groupX // cyxi.
+ by rewrite -def_yp -mulgA mulKg.
+ have [p2 | odd_p] := even_prime p_pr; last first.
+ by rewrite -order_dvdn bin2odd ?dvdn_mulr // -oZ order_dvdG in not_cyxi.
+ have def_yxi: [~ y, x ^+ i] = x ^+ r.
+ have:= Zyxj i; rewrite /Z cycle_traject orderE oZ p2 !inE mulg1.
+ by case/pred2P=> // cyxi; rewrite cyxi p2 eqxx in not_cyxi.
+ apply/existsP; exists (x, x ^+ (i + r %/ 2) * y); rewrite /= !xpair_eqE.
+ rewrite defG x_xjy -order_dvdn ox dvdnn !eqxx andbT /=.
+ rewrite expMg_Rmul /commute ?(centsP cGZ _ (Zyxj _)) ?groupX // def_yxj.
+ rewrite -expgM mulnDl addnC !expgD (expgM x i) -def_yp mulgKV.
+ rewrite -def_yxj def_yxi p2 mulgA -expgD in n_gt23 *.
+ rewrite -expg_mod_order ox /q /r p2 -(subnKC n_gt23) mulnC !expnS mulKn //.
+ rewrite addnn -mul2n modnn mul1g -order_dvdn dvdn_mulr //.
+ by rewrite -p2 -oZ order_dvdG.
+have{i def_yp p_i} Zy2: y ^+ 2 \in Z.
+ rewrite defZ (OhmE _ pX) -groupV -p2 def_yp mem_gen // !inE groupX //= p2.
+ rewrite expgS -{2}def_yp -(mulKg y y) -conjgE -conjXg -conjVg def_yp conjXg.
+ rewrite -expgMn //; last by apply: (centsP cXX); rewrite ?memJ_norm.
+ by rewrite -order_dvdn (dvdn_trans (order_dvdG Z_xxy)) ?oZ.
+rewrite !cycle_traject !orderE oZ p2 !inE !mulg1 /= in Z_xxy Zy2 *.
+rewrite -eq_invg_mul eq_sym -[r]prednK // expgS (inj_eq (mulgI _)) in Z_xxy.
+case/pred2P: Z_xxy => xy; last first.
+ suffices ->: cG = SemiDihedral by []; apply/semidihedral_classP.
+ have n_gt3: n > 3.
+ case: ltngtP notXy => // [|n3]; first by rewrite ltnNge n_gt2.
+ rewrite -scXG inE Gy defX cent_cycle; case/cent1P; red.
+ by rewrite (conjgC x) xy /r p2 n3.
+ exists n => //; rewrite isogEcard card_semidihedral // oG p2 leqnn andbT.
+ rewrite Grp_semidihedral //; apply/existsP=> /=.
+ case/pred2P: Zy2 => y2; [exists (x, y) | exists (x, x * y)].
+ by rewrite /= -{1}[y]mul1g (defG 0%N) y2 xy -p2 -/q -ox expg_order.
+ rewrite /= (defG 1%N) conjgM {2}/conjg mulKg -p2 -/q -ox expg_order -xy.
+ rewrite !xpair_eqE !eqxx /= andbT p2 expgS {2}(conjgC x) xy mulgA -(mulgA x).
+ rewrite [y * y]y2 -expgS -expgD addSnnS prednK // addnn -mul2n -p2 def_pr.
+ by rewrite -ox expg_order.
+case/pred2P: Zy2 => y2.
+ suffices ->: cG = Dihedral by []; apply/dihedral_classP.
+ exists n => //; rewrite isogEcard card_2dihedral // oG p2 leqnn andbT.
+ rewrite Grp_2dihedral //; apply/existsP; exists (x, y) => /=.
+ by rewrite /= -{1}[y]mul1g (defG 0%N) y2 xy -p2 -/q -ox expg_order.
+suffices ->: cG = Quaternion by []; apply/quaternion_classP.
+exists n => //; rewrite isogEcard card_quaternion // oG p2 leqnn andbT.
+rewrite Grp_quaternion //; apply/existsP; exists (x, y) => /=.
+by rewrite /= -{1}[y]mul1g (defG 0%N) y2 xy -p2 -/q -ox expg_order.
+Qed.
+
+(* This is Aschbacher (23.5) *)
+Lemma cyclic_SCN gT p (G U : {group gT}) :
+ p.-group G -> U \in 'SCN(G) -> ~~ abelian G -> cyclic U ->
+ [/\ p = 2, #|G : U| = 2 & extremal2 G]
+\/ exists M : {group gT},
+ [/\ M :=: 'C_G('Mho^1(U)), #|M : U| = p, extremal_class M = ModularGroup,
+ 'Ohm_1(M)%G \in 'E_p^2(G) & 'Ohm_1(M) \char G].
+Proof.
+move=> pG /SCN_P[nsUG scUG] not_cGG cycU; have [sUG nUG] := andP nsUG.
+have [cUU pU] := (cyclic_abelian cycU, pgroupS sUG pG).
+have ltUG: ~~ (G \subset U).
+ by apply: contra not_cGG => sGU; exact: abelianS cUU.
+have ntU: U :!=: 1.
+ by apply: contra ltUG; move/eqP=> U1; rewrite -(setIidPl (cents1 G)) -U1 scUG.
+have [p_pr _ [n oU]] := pgroup_pdiv pU ntU.
+have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1.
+have [u defU] := cyclicP cycU; have Uu: u \in U by rewrite defU cycle_id.
+have Gu := subsetP sUG u Uu; have p_u := mem_p_elt pG Gu.
+have defU1: 'Mho^1(U) = <[u ^+ p]> by rewrite defU (Mho_p_cycle _ p_u).
+have modM1 (M : {group gT}):
+ [/\ U \subset M, #|M : U| = p & extremal_class M = ModularGroup] ->
+ M :=: 'C_M('Mho^1(U)) /\ 'Ohm_1(M)%G \in 'E_p^2(M).
+- case=> sUM iUM /modular_group_classP[q q_pr {n oU}[n n_gt23 isoM]].
+ have n_gt2: n > 2 by exact: leq_trans (leq_addl _ _) n_gt23.
+ have def_n: n = (n - 3).+3 by rewrite -{1}(subnKC n_gt2).
+ have oM: #|M| = (q ^ n)%N by rewrite (card_isog isoM) card_modular_group.
+ have pM: q.-group M by rewrite /pgroup oM pnat_exp pnat_id.
+ have def_q: q = p; last rewrite {q q_pr}def_q in oM pM isoM n_gt23.
+ by apply/eqP; rewrite eq_sym [p == q](pgroupP pM) // -iUM dvdn_indexg.
+ have [[x y] genM modM] := generators_modular_group p_pr n_gt2 isoM.
+ case/modular_group_structure: genM => // _ [defZ _ oZ] _ defMho.
+ have ->: 'Mho^1(U) = 'Z(M).
+ apply/eqP; rewrite eqEcard oZ defZ -(defMho 1%N) ?MhoS //= defU1 -orderE.
+ suff ou: #[u] = (p * p ^ n.-2)%N by rewrite orderXdiv ou ?dvdn_mulr ?mulKn.
+ by rewrite orderE -defU -(divg_indexS sUM) iUM oM def_n mulKn.
+ case: eqP => [[p2 n3] | _ defOhm]; first by rewrite p2 n3 in n_gt23.
+ have{defOhm} [|defM1 oM1] := defOhm 1%N; first by rewrite def_n.
+ split; rewrite ?(setIidPl _) //; first by rewrite centsC subsetIr.
+ rewrite inE oM1 pfactorK // andbT inE Ohm_sub abelem_Ohm1 //.
+ exact: (card_p2group_abelian p_pr oM1).
+have ou: #[u] = (p ^ n.+1)%N by rewrite defU in oU.
+pose Gs := G / U; have pGs: p.-group Gs by rewrite quotient_pgroup.
+have ntGs: Gs != 1 by rewrite -subG1 quotient_sub1.
+have [_ _ [[|k] oGs]] := pgroup_pdiv pGs ntGs.
+ have iUG: #|G : U| = p by rewrite -card_quotient ?oGs.
+ case: (predU1P (maximal_cycle_extremal _ _ _ _ iUG)) => // [modG | ext2G].
+ by right; exists G; case: (modM1 G) => // <- ->; rewrite Ohm_char.
+ by left; case: eqP ext2G => // <-.
+pose M := 'C_G('Mho^1(U)); right; exists [group of M].
+have sMG: M \subset G by exact: subsetIl.
+have [pM nUM] := (pgroupS sMG pG, subset_trans sMG nUG).
+have sUM: U \subset M by rewrite subsetI sUG sub_abelian_cent ?Mho_sub.
+pose A := Aut U; have cAA: abelian A by rewrite Aut_cyclic_abelian.
+have sylAp: p.-Sylow(A) 'O_p(A) by rewrite nilpotent_pcore_Hall ?abelian_nil.
+have [f [injf sfGsA fG]]: exists f : {morphism Gs >-> {perm gT}},
+ [/\ 'injm f, f @* Gs \subset A & {in G, forall y, f (coset U y) u = u ^ y}].
+- have [] := first_isom_loc [morphism of conj_aut U] nUG.
+ rewrite ker_conj_aut scUG /= -/Gs => f injf im_f.
+ exists f; rewrite im_f ?Aut_conj_aut //.
+ split=> // y Gy; have nUy := subsetP nUG y Gy.
+ suffices ->: f (coset U y) = conj_aut U y by rewrite norm_conj_autE.
+ by apply: set1_inj; rewrite -!morphim_set1 ?mem_quotient // im_f ?sub1set.
+have cGsGs: abelian Gs by rewrite -(injm_abelian injf) // (abelianS sfGsA).
+have p_fGs: p.-group (f @* Gs) by rewrite morphim_pgroup.
+have sfGsAp: f @* Gs \subset 'O_p(A) by rewrite (sub_Hall_pcore sylAp).
+have [a [fGa oa au n_gt01 cycGs]]: exists a,
+ [/\ a \in f @* Gs, #[a] = p, a u = u ^+ (p ^ n).+1, (p == 2) + 1 <= n
+ & cyclic Gs \/ p = 2 /\ (exists2 c, c \in f @* Gs & c u = u^-1)].
+- have [m [[def_m _ _ _ _] _]] := cyclic_pgroup_Aut_structure pU cycU ntU.
+ have ->: logn p #|U| = n.+1 by rewrite oU pfactorK.
+ rewrite /= -/A; case: posnP => [_ defA | n_gt0 [c [Ac oc m_c defA]]].
+ have:= cardSg sfGsAp; rewrite (card_Hall sylAp) /= -/A defA card_injm //.
+ by rewrite oGs (part_p'nat (pcore_pgroup _ _)) pfactor_dvdn // logn1.
+ have [p2 | odd_p] := even_prime p_pr; last first.
+ case: eqP => [-> // | _] in odd_p *; rewrite odd_p in defA.
+ have [[cycA _] _ [a [Aa oa m_a defA1]]] := defA.
+ exists a; rewrite -def_m // oa m_a expg_znat //.
+ split=> //; last by left; rewrite -(injm_cyclic injf) ?(cyclicS sfGsA).
+ have: f @* Gs != 1 by rewrite morphim_injm_eq1.
+ rewrite -cycle_subG; apply: contraR => not_sfGs_a.
+ by rewrite -(setIidPl sfGsAp) TI_Ohm1 // defA1 setIC prime_TIg -?orderE ?oa.
+ do [rewrite {1}p2 /= eqn_leq n_gt0; case: leqP => /= [_ | n_gt1]] in defA.
+ have:= cardSg sfGsAp; rewrite (card_Hall sylAp) /= -/A defA -orderE oc p2.
+ by rewrite card_injm // oGs p2 pfactor_dvdn // p_part.
+ have{defA} [s [As os _ defA [a [Aa oa m_a _ defA1]]]] := defA; exists a.
+ have fGs_a: a \in f @* Gs.
+ suffices: f @* Gs :&: <[s]> != 1.
+ apply: contraR => not_fGs_a; rewrite TI_Ohm1 // defA1 setIC.
+ by rewrite prime_TIg -?orderE ?oa // cycle_subG.
+ have: (f @* Gs) * <[s]> \subset A by rewrite mulG_subG cycle_subG sfGsA.
+ move/subset_leq_card; apply: contraL; move/eqP; move/TI_cardMg->.
+ rewrite -(dprod_card defA) -ltnNge mulnC -!orderE ltn_pmul2r // oc.
+ by rewrite card_injm // oGs p2 (ltn_exp2l 1%N).
+ rewrite -def_m // oa m_a expg_znat // p2; split=> //.
+ rewrite abelian_rank1_cyclic // (rank_pgroup pGs) //.
+ rewrite -(injm_p_rank injf) // p_rank_abelian 1?morphim_abelian //= p2 -/Gs.
+ case: leqP => [|fGs1_gt1]; [by left | right].
+ split=> //; exists c; last by rewrite -def_m // m_c expg_zneg.
+ have{defA1} defA1: <[a]> \x <[c]> = 'Ohm_1(Aut U).
+ by rewrite -(Ohm_dprod 1 defA) defA1 (@Ohm_p_cycle 1 _ 2) /p_elt oc.
+ have def_fGs1: 'Ohm_1(f @* Gs) = 'Ohm_1(A).
+ apply/eqP; rewrite eqEcard OhmS // -(dprod_card defA1) -!orderE oa oc.
+ by rewrite dvdn_leq ?(@pfactor_dvdn 2 2) ?cardG_gt0.
+ rewrite (subsetP (Ohm_sub 1 _)) // def_fGs1 -cycle_subG.
+ by case/dprodP: defA1 => _ <- _ _; rewrite mulG_subr.
+have n_gt0: n > 0 := leq_trans (leq_addl _ _) n_gt01.
+have [ys Gys _ def_a] := morphimP fGa.
+have oys: #[ys] = p by rewrite -(order_injm injf) // -def_a oa.
+have defMs: M / U = <[ys]>.
+ apply/eqP; rewrite eq_sym eqEcard -orderE oys cycle_subG; apply/andP; split.
+ have [y nUy Gy /= def_ys] := morphimP Gys.
+ rewrite def_ys mem_quotient //= inE Gy defU1 cent_cycle cent1C.
+ rewrite (sameP cent1P commgP) commgEl conjXg -fG //= -def_ys -def_a au.
+ by rewrite -expgM mulSn expgD mulKg -expnSr -ou expg_order.
+ rewrite card_quotient // -(setIidPr sUM) -scUG setIA (setIidPl sMG).
+ rewrite defU cent_cycle index_cent1 -(card_imset _ (mulgI u^-1)) -imset_comp.
+ have <-: #|'Ohm_1(U)| = p.
+ rewrite defU (Ohm_p_cycle 1 p_u) -orderE (orderXexp _ ou) ou pfactorK //.
+ by rewrite subKn.
+ rewrite (OhmE 1 pU) subset_leq_card ?sub_gen //.
+ apply/subsetP=> _ /imsetP[z /setIP[/(subsetP nUG) nUz cU1z] ->].
+ have Uv' := groupVr Uu; have Uuz: u ^ z \in U by rewrite memJ_norm.
+ rewrite !inE groupM // expgMn /commute 1?(centsP cUU u^-1) //= expgVn -conjXg.
+ by rewrite (sameP commgP cent1P) cent1C -cent_cycle -defU1.
+have iUM: #|M : U| = p by rewrite -card_quotient ?defMs.
+have not_cMM: ~~ abelian M.
+ apply: contraL p_pr => cMM; rewrite -iUM -indexgI /= -/M.
+ by rewrite (setIidPl _) ?indexgg // -scUG subsetI sMG sub_abelian_cent.
+have modM: extremal_class M = ModularGroup.
+ have sU1Z: 'Mho^1(U) \subset 'Z(M).
+ by rewrite subsetI (subset_trans (Mho_sub 1 U)) // centsC subsetIr.
+ case: (predU1P (maximal_cycle_extremal _ _ _ _ iUM)) => //=; rewrite -/M.
+ case/andP; move/eqP=> p2 ext2M; rewrite p2 add1n in n_gt01.
+ suffices{sU1Z}: #|'Z(M)| = 2.
+ move/eqP; rewrite eqn_leq leqNgt (leq_trans _ (subset_leq_card sU1Z)) //.
+ by rewrite defU1 -orderE (orderXexp 1 ou) subn1 p2 (ltn_exp2l 1).
+ move: ext2M; rewrite /extremal2 !inE orbC -orbA; case/or3P; move/eqP.
+ - case/semidihedral_classP=> m m_gt3 isoM.
+ have [[x z] genM [oz _]] := generators_semidihedral m_gt3 isoM.
+ by case/semidihedral_structure: genM => // _ _ [].
+ - case/quaternion_classP=> m m_gt2 isoM.
+ have [[x z] genM _] := generators_quaternion m_gt2 isoM.
+ by case/quaternion_structure: genM => // _ _ [].
+ case/dihedral_classP=> m m_gt1 isoM.
+ have [[x z] genM _] := generators_2dihedral m_gt1 isoM.
+ case/dihedral2_structure: genM not_cMM => // _ _ _ _.
+ by case: (m == 2) => [|[]//]; move/abelem_abelian->.
+split=> //.
+ have [//|_] := modM1 [group of M]; rewrite !inE -andbA /=.
+ by case/andP; move/subset_trans->.
+have{cycGs} [cycGs | [p2 [c fGs_c u_c]]] := cycGs.
+ suffices ->: 'Ohm_1(M) = 'Ohm_1(G) by exact: Ohm_char.
+ suffices sG1M: 'Ohm_1(G) \subset M.
+ by apply/eqP; rewrite eqEsubset -{2}(Ohm_id 1 G) !OhmS.
+ rewrite -(quotientSGK _ sUM) ?(subset_trans (Ohm_sub _ G)) //= defMs.
+ suffices ->: <[ys]> = 'Ohm_1(Gs) by rewrite morphim_Ohm.
+ apply/eqP; rewrite eqEcard -orderE cycle_subG /= {1}(OhmE 1 pGs) /=.
+ rewrite mem_gen ?inE ?Gys -?order_dvdn oys //=.
+ rewrite -(part_pnat_id (pgroupS (Ohm_sub _ _) pGs)) p_part (leq_exp2l _ 1) //.
+ by rewrite -p_rank_abelian -?rank_pgroup -?abelian_rank1_cyclic.
+suffices charU1: 'Mho^1(U) \char G^`(1).
+ rewrite (char_trans (Ohm_char _ _)) // subcent_char ?char_refl //.
+ exact: char_trans (der_char 1 G).
+suffices sUiG': 'Mho^1(U) \subset G^`(1).
+ have cycG': cyclic G^`(1) by rewrite (cyclicS _ cycU) // der1_min.
+ by case/cyclicP: cycG' sUiG' => zs ->; exact: cycle_subgroup_char.
+rewrite defU1 cycle_subG p2 -groupV invMg -{2}u_c.
+by case/morphimP: fGs_c => _ _ /morphimP[z _ Gz ->] ->; rewrite fG ?mem_commg.
+Qed.
+
+(* This is Aschbacher, exercise (8.4) *)
+Lemma normal_rank1_structure gT p (G : {group gT}) :
+ p.-group G -> (forall X : {group gT}, X <| G -> abelian X -> cyclic X) ->
+ cyclic G \/ [&& p == 2, extremal2 G & (#|G| >= 16) || (G \isog 'Q_8)].
+Proof.
+move=> pG dn_G_1.
+have [cGG | not_cGG] := boolP (abelian G); first by left; rewrite dn_G_1.
+have [X maxX]: {X | [max X | X <| G & abelian X]}.
+ by apply: ex_maxgroup; exists 1%G; rewrite normal1 abelian1.
+have cycX: cyclic X by rewrite dn_G_1; case/andP: (maxgroupp maxX).
+have scX: X \in 'SCN(G) := max_SCN pG maxX.
+have [[p2 _ cG] | [M [_ _ _]]] := cyclic_SCN pG scX not_cGG cycX; last first.
+ rewrite 2!inE -andbA; case/and3P=> sEG abelE dimE_2 charE.
+ have:= dn_G_1 _ (char_normal charE) (abelem_abelian abelE).
+ by rewrite (abelem_cyclic abelE) (eqP dimE_2).
+have [n oG] := p_natP pG; right; rewrite p2 cG /= in oG *.
+rewrite oG (@leq_exp2l 2 4) //.
+rewrite /extremal2 /extremal_class oG pfactorKpdiv // in cG.
+case: andP cG => [[n_gt1 isoG] _ | _]; last first.
+ by rewrite leq_eqVlt; case: (3 < n); case: eqP => //= <-; do 2?case: ifP.
+have [[x y] genG _] := generators_2dihedral n_gt1 isoG.
+have [_ _ _ [_ _ maxG]] := dihedral2_structure n_gt1 genG isoG.
+rewrite 2!ltn_neqAle n_gt1 !(eq_sym _ n).
+case: eqP => [_ abelG| _]; first by rewrite (abelem_abelian abelG) in not_cGG.
+case: eqP => // -> [_ _ isoY _ _]; set Y := <<_>> in isoY.
+have nxYG: Y <| G by rewrite (p_maximal_normal pG) // maxG !inE eqxx orbT.
+have [// | [u v] genY _] := generators_2dihedral _ isoY.
+case/dihedral2_structure: (genY) => //= _ _ _ _ abelY.
+have:= dn_G_1 _ nxYG (abelem_abelian abelY).
+by rewrite (abelem_cyclic abelY); case: genY => ->.
+Qed.
+
+(* Replacement for Section 4 proof. *)
+Lemma odd_pgroup_rank1_cyclic gT p (G : {group gT}) :
+ p.-group G -> odd #|G| -> cyclic G = ('r_p(G) <= 1).
+Proof.
+move=> pG oddG; rewrite -rank_pgroup //; apply/idP/idP=> [cycG | dimG1].
+ by rewrite -abelian_rank1_cyclic ?cyclic_abelian.
+have [X nsXG cXX|//|] := normal_rank1_structure pG; last first.
+ by rewrite (negPf (odd_not_extremal2 oddG)) andbF.
+by rewrite abelian_rank1_cyclic // (leq_trans (rankS (normal_sub nsXG))).
+Qed.
+
+(* This is the second part of Aschbacher, exercise (8.4). *)
+Lemma prime_Ohm1P gT p (G : {group gT}) :
+ p.-group G -> G :!=: 1 ->
+ reflect (#|'Ohm_1(G)| = p)
+ (cyclic G || (p == 2) && (extremal_class G == Quaternion)).
+Proof.
+move=> pG ntG; have [p_pr p_dvd_G _] := pgroup_pdiv pG ntG.
+apply: (iffP idP) => [|oG1p].
+ case/orP=> [cycG|]; first exact: Ohm1_cyclic_pgroup_prime.
+ case/andP=> /eqP p2 /eqP/quaternion_classP[n n_gt2 isoG].
+ rewrite p2; have [[x y]] := generators_quaternion n_gt2 isoG.
+ by case/quaternion_structure=> // _ _ [<- oZ _ [->]].
+have [X nsXG cXX|-> //|]:= normal_rank1_structure pG.
+ have [sXG _] := andP nsXG; have pX := pgroupS sXG pG.
+ rewrite abelian_rank1_cyclic // (rank_pgroup pX) p_rank_abelian //.
+ rewrite -{2}(pfactorK 1 p_pr) -{3}oG1p dvdn_leq_log ?cardG_gt0 //.
+ by rewrite cardSg ?OhmS.
+case/and3P=> /eqP p2; rewrite p2 (orbC (cyclic G)) /extremal2.
+case cG: (extremal_class G) => //; case: notF.
+ case/dihedral_classP: cG => n n_gt1 isoG.
+ have [[x y] genG _] := generators_2dihedral n_gt1 isoG.
+ have [oG _ _ _] := genG; case/dihedral2_structure: genG => // _ _ [defG1 _] _.
+ by case/idPn: n_gt1; rewrite -(@ltn_exp2l 2) // -oG -defG1 oG1p p2.
+case/semidihedral_classP: cG => n n_gt3 isoG.
+have [[x y] genG [oy _]] := generators_semidihedral n_gt3 isoG.
+case/semidihedral_structure: genG => // _ _ [_ _ [defG1 _] _] _ [isoG1 _ _].
+case/idPn: (n_gt3); rewrite -(ltn_predK n_gt3) ltnS -leqNgt -(@leq_exp2l 2) //.
+rewrite -card_2dihedral //; last by rewrite -(subnKC n_gt3).
+by rewrite -(card_isog isoG1) /= -defG1 oG1p p2.
+Qed.
+
+(* This is Aschbacher (23.9) *)
+Theorem symplectic_type_group_structure gT p (G : {group gT}) :
+ p.-group G -> (forall X : {group gT}, X \char G -> abelian X -> cyclic X) ->
+ exists2 E : {group gT}, E :=: 1 \/ extraspecial E
+ & exists R : {group gT},
+ [/\ cyclic R \/ [/\ p = 2, extremal2 R & #|R| >= 16],
+ E \* R = G
+ & E :&: R = 'Z(E)].
+Proof.
+move=> pG sympG; have [H [charH]] := Thompson_critical pG.
+have sHG := char_sub charH; have pH := pgroupS sHG pG.
+set U := 'Z(H) => sPhiH_U sHG_U defU; set Z := 'Ohm_1(U).
+have sZU: Z \subset U by rewrite Ohm_sub.
+have charU: U \char G := char_trans (center_char H) charH.
+have cUU: abelian U := center_abelian H.
+have cycU: cyclic U by exact: sympG.
+have pU: p.-group U := pgroupS (char_sub charU) pG.
+have cHU: U \subset 'C(H) by rewrite subsetIr.
+have cHsHs: abelian (H / Z).
+ rewrite sub_der1_abelian //= (OhmE _ pU) genS //= -/U.
+ apply/subsetP=> _ /imset2P[h k Hh Hk ->].
+ have Uhk: [~ h, k] \in U by rewrite (subsetP sHG_U) ?mem_commg ?(subsetP sHG).
+ rewrite inE Uhk inE -commXg; last by red; rewrite -(centsP cHU).
+ apply/commgP; red; rewrite (centsP cHU) // (subsetP sPhiH_U) //.
+ by rewrite (Phi_joing pH) mem_gen // inE orbC (Mho_p_elt 1) ?(mem_p_elt pH).
+have nsZH: Z <| H by rewrite sub_center_normal.
+have [K /=] := inv_quotientS nsZH (Ohm_sub 1 (H / Z)); fold Z => defKs sZK sKH.
+have nsZK: Z <| K := normalS sZK sKH nsZH; have [_ nZK] := andP nsZK.
+have abelKs: p.-abelem (K / Z) by rewrite -defKs Ohm1_abelem ?quotient_pgroup.
+have charK: K \char G.
+ have charZ: Z \char H := char_trans (Ohm_char _ _) (center_char H).
+ rewrite (char_trans _ charH) // (char_from_quotient nsZK) //.
+ by rewrite -defKs Ohm_char.
+have cycZK: cyclic 'Z(K).
+ by rewrite sympG ?center_abelian ?(char_trans (center_char _)).
+have [cKK | not_cKK] := orP (orbN (abelian K)).
+ have defH: U = H.
+ apply: center_idP; apply: cyclic_factor_abelian (Ohm_sub 1 _) _.
+ rewrite /= -/Z abelian_rank1_cyclic //.
+ have cKsKs: abelian (K / Z) by rewrite -defKs (abelianS (Ohm_sub 1 _)).
+ have cycK: cyclic K by rewrite -(center_idP cKK).
+ by rewrite -rank_Ohm1 defKs -abelian_rank1_cyclic ?quotient_cyclic.
+ have scH: H \in 'SCN(G) by apply/SCN_P; rewrite defU char_normal.
+ have [cGG | not_cGG] := orP (orbN (abelian G)).
+ exists 1%G; [by left | exists G; rewrite cprod1g (setIidPl _) ?sub1G //].
+ by split; first left; rewrite ?center1 // sympG ?char_refl.
+ have cycH: cyclic H by rewrite -{}defH.
+ have [[p2 _ cG2]|[M [_ _ _]]] := cyclic_SCN pG scH not_cGG cycH; last first.
+ do 2![case/setIdP] => _ abelE dimE_2 charE.
+ have:= sympG _ charE (abelem_abelian abelE).
+ by rewrite (abelem_cyclic abelE) (eqP dimE_2).
+ have [n oG] := p_natP pG; rewrite p2 in oG.
+ have [n_gt3 | n_le3] := ltnP 3 n.
+ exists 1%G; [by left | exists G; rewrite cprod1g (setIidPl _) ?sub1G //].
+ by split; first right; rewrite ?center1 // oG (@leq_exp2l 2 4).
+ have esG: extraspecial G.
+ by apply: (p3group_extraspecial pG); rewrite // p2 oG pfactorK.
+ exists G; [by right | exists ('Z(G))%G; rewrite cprod_center_id setIA setIid].
+ by split=> //; left; rewrite prime_cyclic; case: esG.
+have ntK: K :!=: 1 by apply: contra not_cKK; move/eqP->; exact: abelian1.
+have [p_pr _ _] := pgroup_pdiv (pgroupS sKH pH) ntK.
+have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1.
+have oZ: #|Z| = p.
+ apply: Ohm1_cyclic_pgroup_prime => //=; apply: contra ntK; move/eqP.
+ by move/(trivg_center_pgroup pH)=> GH; rewrite -subG1 -GH.
+have sZ_ZK: Z \subset 'Z(K).
+ by rewrite subsetI sZK (subset_trans (Ohm_sub _ _ )) // subIset ?centS ?orbT.
+have sZsKs: 'Z(K) / Z \subset K / Z by rewrite quotientS ?center_sub.
+have [Es /= splitKs] := abelem_split_dprod abelKs sZsKs.
+have [_ /= defEsZs cEsZs tiEsZs] := dprodP splitKs.
+have sEsKs: Es \subset K / Z by rewrite -defEsZs mulG_subr.
+have [E defEs sZE sEK] := inv_quotientS nsZK sEsKs; rewrite /= -/Z in defEs sZE.
+have [nZE nZ_ZK] := (subset_trans sEK nZK, subset_trans (center_sub K) nZK).
+have defK: 'Z(K) * E = K.
+ rewrite -(mulSGid sZ_ZK) -mulgA -quotientK ?mul_subG ?quotientMl //.
+ by rewrite -defEs defEsZs quotientGK.
+have defZE: 'Z(E) = Z.
+ have cEZK: 'Z(K) \subset 'C(E) by rewrite subIset // orbC centS.
+ have cE_Z: E \subset 'C(Z) by rewrite centsC (subset_trans sZ_ZK).
+ apply/eqP; rewrite eqEsubset andbC subsetI sZE centsC cE_Z /=.
+ rewrite -quotient_sub1 ?subIset ?nZE //= -/Z -tiEsZs subsetI defEs.
+ rewrite !quotientS ?center_sub //= subsetI subIset ?sEK //=.
+ by rewrite -defK centM setSI // centsC.
+have sEH := subset_trans sEK sKH; have pE := pgroupS sEH pH.
+have esE: extraspecial E.
+ split; last by rewrite defZE oZ.
+ have sPhiZ: 'Phi(E) \subset Z.
+ rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) ?(quotient_Phi pE) //.
+ rewrite subG1 (trivg_Phi (quotient_pgroup _ pE)) /= -defEs.
+ by rewrite (abelemS sEsKs) //= -defKs Ohm1_abelem ?quotient_pgroup.
+ have sE'Phi: E^`(1) \subset 'Phi(E) by rewrite (Phi_joing pE) joing_subl.
+ have ntE': E^`(1) != 1.
+ rewrite (sameP eqP commG1P) -abelianE; apply: contra not_cKK => cEE.
+ by rewrite -defK mulGSid ?center_abelian // -(center_idP cEE) defZE.
+ have defE': E^`(1) = Z.
+ apply/eqP; rewrite eqEcard (subset_trans sE'Phi) //= oZ.
+ have [_ _ [n ->]] := pgroup_pdiv (pgroupS (der_sub _ _) pE) ntE'.
+ by rewrite (leq_exp2l 1) ?prime_gt1.
+ by split; rewrite defZE //; apply/eqP; rewrite eqEsubset sPhiZ -defE'.
+have [spE _] := esE; have [defPhiE defE'] := spE.
+have{defE'} sEG_E': [~: E, G] \subset E^`(1).
+ rewrite defE' defZE /Z (OhmE _ pU) commGC genS //.
+ apply/subsetP=> _ /imset2P[g e Gg Ee ->].
+ have He: e \in H by rewrite (subsetP sKH) ?(subsetP sEK).
+ have Uge: [~ g, e] \in U by rewrite (subsetP sHG_U) ?mem_commg.
+ rewrite inE Uge inE -commgX; last by red; rewrite -(centsP cHU).
+ have sZ_ZG: Z \subset 'Z(G).
+ have charZ: Z \char G := char_trans (Ohm_char _ _) charU.
+ move/implyP: (meet_center_nil (pgroup_nil pG) (char_normal charZ)).
+ rewrite -cardG_gt1 oZ prime_gt1 //=; apply: contraR => not_sZ_ZG.
+ by rewrite prime_TIg ?oZ.
+ have: e ^+ p \in 'Z(G).
+ rewrite (subsetP sZ_ZG) // -defZE -defPhiE (Phi_joing pE) mem_gen //.
+ by rewrite inE orbC (Mho_p_elt 1) ?(mem_p_elt pE).
+ by case/setIP=> _ /centP cGep; apply/commgP; red; rewrite cGep.
+have sEG: E \subset G := subset_trans sEK (char_sub charK).
+set R := 'C_G(E).
+have{sEG_E'} defG: E \* R = G by exact: (critical_extraspecial pG).
+have [_ defER cRE] := cprodP defG.
+have defH: E \* 'C_H(E) = H by rewrite -(setIidPr sHG) setIAC (cprod_modl defG).
+have{defH} [_ defH cRH_E] := cprodP defH.
+have cRH_RH: abelian 'C_H(E).
+ have sZ_ZRH: Z \subset 'Z('C_H(E)).
+ rewrite subsetI -{1}defZE setSI //= (subset_trans sZU) // centsC.
+ by rewrite subIset // centsC cHU.
+ rewrite (cyclic_factor_abelian sZ_ZRH) //= -/Z.
+ have defHs: Es \x ('C_H(E) / Z) = H / Z.
+ rewrite defEs dprodE ?quotient_cents // -?quotientMl ?defH -?quotientGI //=.
+ by rewrite setIA (setIidPl sEH) ['C_E(E)]defZE trivg_quotient.
+ have:= Ohm_dprod 1 defHs; rewrite /= defKs (Ohm1_id (abelemS sEsKs abelKs)).
+ rewrite dprodC; case/dprodP=> _ defEsRHs1 cRHs1Es tiRHs1Es.
+ have sRHsHs: 'C_H(E) / Z \subset H / Z by rewrite quotientS ?subsetIl.
+ have cRHsRHs: abelian ('C_H(E) / Z) by exact: abelianS cHsHs.
+ have pHs: p.-group (H / Z) by rewrite quotient_pgroup.
+ rewrite abelian_rank1_cyclic // (rank_pgroup (pgroupS sRHsHs pHs)).
+ rewrite p_rank_abelian // -(leq_add2r (logn p #|Es|)) -lognM ?cardG_gt0 //.
+ rewrite -TI_cardMg // defEsRHs1 /= -defEsZs TI_cardMg ?lognM ?cardG_gt0 //.
+ by rewrite leq_add2r -abelem_cyclic ?(abelemS sZsKs) // quotient_cyclic.
+have{cRH_RH} defRH: 'C_H(E) = U.
+ apply/eqP; rewrite eqEsubset andbC setIS ?centS // subsetI subsetIl /=.
+ by rewrite -{2}defH centM subsetI subsetIr.
+have scUR: 'C_R(U) = U by rewrite -setIA -{1}defRH -centM defH.
+have sUR: U \subset R by rewrite -defRH setSI.
+have tiER: E :&: R = 'Z(E) by rewrite setIA (setIidPl (subset_trans sEH sHG)).
+have [cRR | not_cRR] := boolP (abelian R).
+ exists E; [by right | exists [group of R]; split=> //; left].
+ by rewrite /= -(setIidPl (sub_abelian_cent cRR sUR)) scUR.
+have{scUR} scUR: [group of U] \in 'SCN(R).
+ by apply/SCN_P; rewrite (normalS sUR (subsetIl _ _)) // char_normal.
+have pR: p.-group R := pgroupS (subsetIl _ _) pG.
+have [R_le_3 | R_gt_3] := leqP (logn p #|R|) 3.
+ have esR: extraspecial R := p3group_extraspecial pR not_cRR R_le_3.
+ have esG: extraspecial G := cprod_extraspecial pG defG tiER esE esR.
+ exists G; [by right | exists ('Z(G))%G; rewrite cprod_center_id setIA setIid].
+ by split=> //; left; rewrite prime_cyclic; case: esG.
+have [[p2 _ ext2R] | [M []]] := cyclic_SCN pR scUR not_cRR cycU.
+ exists E; [by right | exists [group of R]; split=> //; right].
+ by rewrite dvdn_leq ?(@pfactor_dvdn 2 4) ?cardG_gt0 // -{2}p2.
+rewrite /= -/R => defM iUM modM _ _; pose N := 'C_G('Mho^1(U)).
+have charZN2: 'Z('Ohm_2(N)) \char G.
+ rewrite (char_trans (center_char _)) // (char_trans (Ohm_char _ _)) //.
+ by rewrite subcent_char ?char_refl // (char_trans (Mho_char _ _)).
+have:= sympG _ charZN2 (center_abelian _).
+rewrite abelian_rank1_cyclic ?center_abelian // leqNgt; case/negP.
+have defN: E \* M = N.
+ rewrite defM (cprod_modl defG) // centsC (subset_trans (Mho_sub 1 _)) //.
+ by rewrite /= -/U -defRH subsetIr.
+case/modular_group_classP: modM => q q_pr [n n_gt23 isoM].
+have{n_gt23} n_gt2 := leq_trans (leq_addl _ _) n_gt23.
+have n_gt1 := ltnW n_gt2; have n_gt0 := ltnW n_gt1.
+have [[x y] genM modM] := generators_modular_group q_pr n_gt2 isoM.
+have{q_pr} defq: q = p; last rewrite {q}defq in genM modM isoM.
+ have: p %| #|M| by rewrite -iUM dvdn_indexg.
+ by have [-> _ _ _] := genM; rewrite Euclid_dvdX // dvdn_prime2 //; case: eqP.
+have [oM Mx ox X'y] := genM; have [My _] := setDP X'y; have [oy _] := modM.
+have [sUM sMR]: U \subset M /\ M \subset R.
+ by rewrite defM subsetI sUR subsetIl centsC (subset_trans (Mho_sub _ _)).
+have oU1: #|'Mho^1(U)| = (p ^ n.-2)%N.
+ have oU: #|U| = (p ^ n.-1)%N.
+ by rewrite -(divg_indexS sUM) iUM oM -subn1 expnB.
+ case/cyclicP: cycU pU oU => u -> p_u ou.
+ by rewrite (Mho_p_cycle 1 p_u) -orderE (orderXexp 1 ou) subn1.
+have sZU1: Z \subset 'Mho^1(U).
+ rewrite -(cardSg_cyclic cycU) ?Ohm_sub ?Mho_sub // oZ oU1.
+ by rewrite -(subnKC n_gt2) expnS dvdn_mulr.
+case/modular_group_structure: genM => // _ [defZM _ oZM] _ _.
+have:= n_gt2; rewrite leq_eqVlt eq_sym !xpair_eqE andbC.
+case: eqP => [n3 _ _ | _ /= n_gt3 defOhmM].
+ have eqZU1: Z = 'Mho^1(U) by apply/eqP; rewrite eqEcard sZU1 oZ oU1 n3 /=.
+ rewrite (setIidPl _) in defM; first by rewrite -defM oM n3 pfactorK in R_gt_3.
+ by rewrite -eqZU1 subIset ?centS ?orbT.
+have{defOhmM} [|defM2 _] := defOhmM 2; first by rewrite -subn1 ltn_subRL.
+do [set xpn3 := x ^+ _; set X2 := <[_]>] in defM2.
+have oX2: #|X2| = (p ^ 2)%N.
+ by rewrite -orderE (orderXexp _ ox) -{1}(subnKC n_gt2) addSn addnK.
+have sZX2: Z \subset X2.
+ have cycXp: cyclic <[x ^+ p]> := cycle_cyclic _.
+ rewrite -(cardSg_cyclic cycXp) /=; first by rewrite oZ oX2 dvdn_mull.
+ rewrite -defZM subsetI (subset_trans (Ohm_sub _ _)) //=.
+ by rewrite (subset_trans sZU1) // centsC defM subsetIr.
+ by rewrite /xpn3 -subnSK //expnS expgM cycleX.
+have{defM2} [_ /= defM2 cYX2 tiX2Y] := dprodP defM2.
+have{defN} [_ defN cME] := cprodP defN.
+have cEM2: E \subset 'C('Ohm_2(M)).
+ by rewrite centsC (subset_trans _ cME) ?centS ?Ohm_sub.
+have [cEX2 cYE]: X2 \subset 'C(E) /\ E \subset 'C(<[y]>).
+ by apply/andP; rewrite centsC -subsetI -centM defM2.
+have pN: p.-group N := pgroupS (subsetIl _ _) pG.
+have defN2: (E <*> X2) \x <[y]> = 'Ohm_2(N).
+ rewrite dprodE ?centY ?subsetI 1?centsC ?cYE //=; last first.
+ rewrite -cycle_subG in My; rewrite joingC cent_joinEl //= -/X2.
+ rewrite -(setIidPr My) setIA -group_modl ?cycle_subG ?groupX //.
+ by rewrite mulGSid // (subset_trans _ sZX2) // -defZE -tiER setIS.
+ apply/eqP; rewrite cent_joinEr // -mulgA defM2 eqEsubset mulG_subG.
+ rewrite OhmS ?andbT; last by rewrite -defN mulG_subr.
+ have expE: exponent E %| p ^ 2 by rewrite exponent_special ?(pgroupS sEG).
+ rewrite /= (OhmE 2 pN) sub_gen /=; last 1 first.
+ by rewrite subsetI -defN mulG_subl sub_LdivT expE.
+ rewrite -cent_joinEl // -genM_join genS // -defN.
+ apply/subsetP=> ez; case/setIP; case/imset2P=> e z Ee Mz ->{ez}.
+ rewrite inE expgMn; last by red; rewrite -(centsP cME).
+ rewrite (exponentP expE) // mul1g => zp2; rewrite mem_mulg //=.
+ by rewrite (OhmE 2 (pgroupS sMR pR)) mem_gen // !inE Mz.
+have{defN2} defZN2: X2 \x <[y]> = 'Z('Ohm_2(N)).
+ rewrite -[X2](mulSGid sZX2) /= -/Z -defZE -(center_dprod defN2).
+ do 2!rewrite -{1}(center_idP (cycle_abelian _)) -/X2; congr (_ \x _).
+ by case/cprodP: (center_cprod (cprodEY cEX2)).
+have{defZN2} strZN2: \big[dprod/1]_(z <- [:: xpn3; y]) <[z]> = 'Z('Ohm_2(N)).
+ by rewrite unlock /= dprodg1.
+rewrite -size_abelian_type ?center_abelian //.
+have pZN2: p.-group 'Z('Ohm_2(N)) by rewrite (pgroupS _ pN) // subIset ?Ohm_sub.
+rewrite -(perm_eq_size (perm_eq_abelian_type pZN2 strZN2 _)) //= !inE.
+rewrite !(eq_sym 1) -!order_eq1 oy orderE oX2.
+by rewrite (eqn_exp2l 2 0) // (eqn_exp2l 1 0).
+Qed.
+
+End ExtremalTheory.
diff --git a/mathcomp/solvable/finmodule.v b/mathcomp/solvable/finmodule.v
new file mode 100644
index 0000000..e2eae84
--- /dev/null
+++ b/mathcomp/solvable/finmodule.v
@@ -0,0 +1,596 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice.
+Require Import fintype bigop ssralg finset fingroup morphism perm.
+Require Import finalg action gproduct commutator cyclic.
+
+(******************************************************************************)
+(* This file regroups constructions and results that are based on the most *)
+(* primitive version of representation theory -- viewing an abelian group as *)
+(* the additive group of a (finite) Z-module. This includes the Gaschutz *)
+(* splitting and transitivity theorem, from which we will later derive the *)
+(* Schur-Zassenhaus theorem and the elementary abelian special case of *)
+(* Maschke's theorem, the coprime abelian centraliser/commutator trivial *)
+(* intersection theorem, from which we will derive that p-groups under coprime*)
+(* action factor into special groups, and the construction of the transfer *)
+(* homomorphism and its expansion relative to a cycle, from which we derive *)
+(* the Higman Focal Subgroup and the Burnside Normal Complement theorems. *)
+(* The definitions and lemmas for the finite Z-module induced by an abelian *)
+(* are packaged in an auxiliary FiniteModule submodule: they should not be *)
+(* needed much outside this file, which contains all the results that exploit *)
+(* this construction. *)
+(* FiniteModule defines the Z[N(A)]-module associated with a finite abelian *)
+(* abelian group A, given a proof abelA : abelian A) : *)
+(* fmod_of abelA == the type of elements of the module (similar to but *)
+(* distinct from [subg A]). *)
+(* fmod abelA x == the injection of x into fmod_of abelA if x \in A, else 0 *)
+(* fmval u == the projection of u : fmod_of abelA onto A *)
+(* u ^@ x == the action of x \in 'N(A) on u : fmod_of abelA *)
+(* The transfer morphism is be constructed from a morphism f : H >-> rT, and *)
+(* a group G, along with the two assumptions sHG : H \subset G and *)
+(* abfH : abelian (f @* H): *)
+(* transfer sGH abfH == the function gT -> FiniteModule.fmod_of abfH that *)
+(* implements the transfer morphism induced by f on G. *)
+(* The Lemma transfer_indep states that the transfer morphism can be expanded *)
+(* using any transversal of the partition HG := rcosets H G of G. *)
+(* Further, for any g \in G, HG :* <[g]> is also a partition of G (Lemma *)
+(* rcosets_cycle_partition), and for any transversal X of HG :* <[g]> the *)
+(* function r mapping x : gT to rcosets (H :* x) <[g]> is (constructively) a *)
+(* bijection from X to the <[g]>-orbit partition of HG, and Lemma *)
+(* transfer_pcycle_def gives a simplified expansion of the transfer morphism. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope GRing.Theory FinRing.Theory.
+Local Open Scope ring_scope.
+
+Module FiniteModule.
+
+Reserved Notation "u ^@ x" (at level 31, left associativity).
+
+Inductive fmod_of (gT : finGroupType) (A : {group gT}) (abelA : abelian A) :=
+ Fmod x & x \in A.
+
+Bind Scope ring_scope with fmod_of.
+
+Section OneFinMod.
+
+Let f2sub (gT : finGroupType) (A : {group gT}) (abA : abelian A) :=
+ fun u : fmod_of abA => let : Fmod x Ax := u in Subg Ax : FinGroup.arg_sort _.
+Local Coercion f2sub : fmod_of >-> FinGroup.arg_sort.
+
+Variables (gT : finGroupType) (A : {group gT}) (abelA : abelian A).
+Local Notation fmodA := (fmod_of abelA).
+Implicit Types (x y z : gT) (u v w : fmodA).
+
+Let sub2f (s : [subg A]) := Fmod abelA (valP s).
+
+Definition fmval u := val (f2sub u).
+Canonical fmod_subType := [subType for fmval].
+Local Notation valA := (@val _ _ fmod_subType) (only parsing).
+Definition fmod_eqMixin := Eval hnf in [eqMixin of fmodA by <:].
+Canonical fmod_eqType := Eval hnf in EqType fmodA fmod_eqMixin.
+Definition fmod_choiceMixin := [choiceMixin of fmodA by <:].
+Canonical fmod_choiceType := Eval hnf in ChoiceType fmodA fmod_choiceMixin.
+Definition fmod_countMixin := [countMixin of fmodA by <:].
+Canonical fmod_countType := Eval hnf in CountType fmodA fmod_countMixin.
+Canonical fmod_subCountType := Eval hnf in [subCountType of fmodA].
+Definition fmod_finMixin := [finMixin of fmodA by <:].
+Canonical fmod_finType := Eval hnf in FinType fmodA fmod_finMixin.
+Canonical fmod_subFinType := Eval hnf in [subFinType of fmodA].
+
+Definition fmod x := sub2f (subg A x).
+Definition actr u x := if x \in 'N(A) then fmod (fmval u ^ x) else u.
+
+Definition fmod_opp u := sub2f u^-1.
+Definition fmod_add u v := sub2f (u * v).
+
+Fact fmod_add0r : left_id (sub2f 1) fmod_add.
+Proof. move=> u; apply: val_inj; exact: mul1g. Qed.
+
+Fact fmod_addrA : associative fmod_add.
+Proof. move=> u v w; apply: val_inj; exact: mulgA. Qed.
+
+Fact fmod_addNr : left_inverse (sub2f 1) fmod_opp fmod_add.
+Proof. move=> u; apply: val_inj; exact: mulVg. Qed.
+
+Fact fmod_addrC : commutative fmod_add.
+Proof. case=> x Ax [y Ay]; apply: val_inj; exact: (centsP abelA). Qed.
+
+Definition fmod_zmodMixin :=
+ ZmodMixin fmod_addrA fmod_addrC fmod_add0r fmod_addNr.
+Canonical fmod_zmodType := Eval hnf in ZmodType fmodA fmod_zmodMixin.
+Canonical fmod_finZmodType := Eval hnf in [finZmodType of fmodA].
+Canonical fmod_baseFinGroupType :=
+ Eval hnf in [baseFinGroupType of fmodA for +%R].
+Canonical fmod_finGroupType :=
+ Eval hnf in [finGroupType of fmodA for +%R].
+
+Lemma fmodP u : val u \in A. Proof. exact: valP. Qed.
+Lemma fmod_inj : injective fmval. Proof. exact: val_inj. Qed.
+Lemma congr_fmod u v : u = v -> fmval u = fmval v.
+Proof. exact: congr1. Qed.
+
+Lemma fmvalA : {morph valA : x y / x + y >-> (x * y)%g}. Proof. by []. Qed.
+Lemma fmvalN : {morph valA : x / - x >-> x^-1%g}. Proof. by []. Qed.
+Lemma fmval0 : valA 0 = 1%g. Proof. by []. Qed.
+Canonical fmval_morphism := @Morphism _ _ setT fmval (in2W fmvalA).
+
+Definition fmval_sum := big_morph fmval fmvalA fmval0.
+
+Lemma fmvalZ n : {morph valA : x / x *+ n >-> (x ^+ n)%g}.
+Proof. by move=> u; rewrite /= morphX ?inE. Qed.
+
+Lemma fmodKcond x : val (fmod x) = if x \in A then x else 1%g.
+Proof. by rewrite /= /fmval /= val_insubd. Qed.
+Lemma fmodK : {in A, cancel fmod val}. Proof. exact: subgK. Qed.
+Lemma fmvalK : cancel val fmod.
+Proof. by case=> x Ax; apply: val_inj; rewrite /fmod /= sgvalK. Qed.
+Lemma fmod1 : fmod 1 = 0. Proof. by rewrite -fmval0 fmvalK. Qed.
+Lemma fmodM : {in A &, {morph fmod : x y / (x * y)%g >-> x + y}}.
+Proof. by move=> x y Ax Ay /=; apply: val_inj; rewrite /fmod morphM. Qed.
+Canonical fmod_morphism := Morphism fmodM.
+Lemma fmodX n : {in A, {morph fmod : x / (x ^+ n)%g >-> x *+ n}}.
+Proof. exact: morphX. Qed.
+Lemma fmodV : {morph fmod : x / x^-1%g >-> - x}.
+Proof.
+move=> x; apply: val_inj; rewrite fmvalN !fmodKcond groupV.
+by case: (x \in A); rewrite ?invg1.
+Qed.
+
+Lemma injm_fmod : 'injm fmod.
+Proof.
+apply/injmP=> x y Ax Ay []; move/val_inj; exact: (injmP (injm_subg A)).
+Qed.
+
+Notation "u ^@ x" := (actr u x) : ring_scope.
+
+Lemma fmvalJcond u x :
+ val (u ^@ x) = if x \in 'N(A) then val u ^ x else val u.
+Proof. by case: ifP => Nx; rewrite /actr Nx ?fmodK // memJ_norm ?fmodP. Qed.
+
+Lemma fmvalJ u x : x \in 'N(A) -> val (u ^@ x) = val u ^ x.
+Proof. by move=> Nx; rewrite fmvalJcond Nx. Qed.
+
+Lemma fmodJ x y : y \in 'N(A) -> fmod (x ^ y) = fmod x ^@ y.
+Proof.
+move=> Ny; apply: val_inj; rewrite fmvalJ ?fmodKcond ?memJ_norm //.
+by case: ifP => // _; rewrite conj1g.
+Qed.
+
+Fact actr_is_action : is_action 'N(A) actr.
+Proof.
+split=> [a u v eq_uv_a | u a b Na Nb].
+ case Na: (a \in 'N(A)); last by rewrite /actr Na in eq_uv_a.
+ by apply: val_inj; apply: (conjg_inj a); rewrite -!fmvalJ ?eq_uv_a.
+by apply: val_inj; rewrite !fmvalJ ?groupM ?conjgM.
+Qed.
+
+Canonical actr_action := Action actr_is_action.
+Notation "''M'" := actr_action (at level 8) : action_scope.
+
+Lemma act0r x : 0 ^@ x = 0.
+Proof. by rewrite /actr conj1g morph1 if_same. Qed.
+
+Lemma actAr x : {morph actr^~ x : u v / u + v}.
+Proof.
+by move=> u v; apply: val_inj; rewrite !(fmvalA, fmvalJcond) conjMg; case: ifP.
+Qed.
+
+Definition actr_sum x := big_morph _ (actAr x) (act0r x).
+
+Lemma actNr x : {morph actr^~ x : u / - u}.
+Proof. by move=> u; apply: (addrI (u ^@ x)); rewrite -actAr !subrr act0r. Qed.
+
+Lemma actZr x n : {morph actr^~ x : u / u *+ n}.
+Proof.
+by move=> u; elim: n => [|n IHn]; rewrite ?act0r // !mulrS actAr IHn.
+Qed.
+
+Fact actr_is_groupAction : is_groupAction setT 'M.
+Proof.
+move=> a Na /=; rewrite inE; apply/andP; split.
+ by apply/subsetP=> u _; rewrite inE.
+by apply/morphicP=> u v _ _; rewrite !permE /= actAr.
+Qed.
+
+Canonical actr_groupAction := GroupAction actr_is_groupAction.
+Notation "''M'" := actr_groupAction (at level 8) : groupAction_scope.
+
+Lemma actr1 u : u ^@ 1 = u.
+Proof. exact: act1. Qed.
+
+Lemma actrM : {in 'N(A) &, forall x y u, u ^@ (x * y) = u ^@ x ^@ y}.
+Proof.
+by move=> x y Nx Ny /= u; apply: val_inj; rewrite !fmvalJ ?conjgM ?groupM.
+Qed.
+
+Lemma actrK x : cancel (actr^~ x) (actr^~ x^-1%g).
+Proof.
+move=> u; apply: val_inj; rewrite !fmvalJcond groupV.
+by case: ifP => -> //; rewrite conjgK.
+Qed.
+
+Lemma actrKV x : cancel (actr^~ x^-1%g) (actr^~ x).
+Proof. by move=> u; rewrite /= -{2}(invgK x) actrK. Qed.
+
+End OneFinMod.
+
+Bind Scope ring_scope with fmod_of.
+Prenex Implicits fmval fmod actr.
+Notation "u ^@ x" := (actr u x) : ring_scope.
+Notation "''M'" := actr_action (at level 8) : action_scope.
+Notation "''M'" := actr_groupAction : groupAction_scope.
+
+End FiniteModule.
+
+Canonical FiniteModule.fmod_subType.
+Canonical FiniteModule.fmod_eqType.
+Canonical FiniteModule.fmod_choiceType.
+Canonical FiniteModule.fmod_countType.
+Canonical FiniteModule.fmod_finType.
+Canonical FiniteModule.fmod_subCountType.
+Canonical FiniteModule.fmod_subFinType.
+Canonical FiniteModule.fmod_zmodType.
+Canonical FiniteModule.fmod_finZmodType.
+Canonical FiniteModule.fmod_baseFinGroupType.
+Canonical FiniteModule.fmod_finGroupType.
+
+(* Still allow ring notations, but give priority to groups now. *)
+Import FiniteModule GroupScope.
+
+Section Gaschutz.
+
+Variables (gT : finGroupType) (G H P : {group gT}).
+Implicit Types K L : {group gT}.
+
+Hypotheses (nsHG : H <| G) (sHP : H \subset P) (sPG : P \subset G).
+Hypotheses (abelH : abelian H) (coHiPG : coprime #|H| #|G : P|).
+
+Let sHG := normal_sub nsHG.
+Let nHG := subsetP (normal_norm nsHG).
+
+Let m := (expg_invn H #|G : P|).
+
+Implicit Types a b : fmod_of abelH.
+Local Notation fmod := (fmod abelH).
+
+Theorem Gaschutz_split : [splits G, over H] = [splits P, over H].
+Proof.
+apply/splitsP/splitsP=> [[K /complP[tiHK eqHK]] | [Q /complP[tiHQ eqHQ]]].
+ exists (K :&: P)%G; rewrite inE setICA (setIidPl sHP) setIC tiHK eqxx.
+ by rewrite group_modl // eqHK (sameP eqP setIidPr).
+have sQP: Q \subset P by rewrite -eqHQ mulG_subr.
+pose rP x := repr (P :* x); pose pP x := x * (rP x)^-1.
+have PpP x: pP x \in P by rewrite -mem_rcoset rcoset_repr rcoset_refl.
+have rPmul x y: x \in P -> rP (x * y) = rP y.
+ by move=> Px; rewrite /rP rcosetM rcoset_id.
+pose pQ x := remgr H Q x; pose rH x := pQ (pP x) * rP x.
+have pQhq: {in H & Q, forall h q, pQ (h * q) = q} by exact: remgrMid.
+have pQmul: {in P &, {morph pQ : x y / x * y}}.
+ apply: remgrM; [exact/complP | exact: normalS (nsHG)].
+have HrH x: rH x \in H :* x.
+ by rewrite rcoset_sym mem_rcoset invMg mulgA mem_divgr // eqHQ PpP.
+have GrH x: x \in G -> rH x \in G.
+ move=> Gx; case/rcosetP: (HrH x) => y Hy ->.
+ by rewrite groupM // (subsetP sHG).
+have rH_Pmul x y: x \in P -> rH (x * y) = pQ x * rH y.
+ by move=> Px; rewrite /rH mulgA -pQmul; first by rewrite /pP rPmul ?mulgA.
+have rH_Hmul h y: h \in H -> rH (h * y) = rH y.
+ by move=> Hh; rewrite rH_Pmul ?(subsetP sHP) // -(mulg1 h) pQhq ?mul1g.
+pose mu x y := fmod ((rH x * rH y)^-1 * rH (x * y)).
+pose nu y := (\sum_(Px in rcosets P G) mu (repr Px) y)%R.
+have rHmul: {in G &, forall x y, rH (x * y) = rH x * rH y * val (mu x y)}.
+ move=> x y Gx Gy; rewrite /= fmodK ?mulKVg // -mem_lcoset lcoset_sym.
+ rewrite -norm_rlcoset; last by rewrite nHG ?GrH ?groupM.
+ by rewrite (rcoset_transl (HrH _)) -rcoset_mul ?nHG ?GrH // mem_mulg.
+have actrH a x: x \in G -> (a ^@ rH x = a ^@ x)%R.
+ move=> Gx; apply: val_inj; rewrite /= !fmvalJ ?nHG ?GrH //.
+ case/rcosetP: (HrH x) => b /(fmodK abelH) <- ->; rewrite conjgM.
+ by congr (_ ^ _); rewrite conjgE -fmvalN -!fmvalA (addrC a) addKr.
+have mu_Pmul x y z: x \in P -> mu (x * y) z = mu y z.
+ move=> Px; congr fmod; rewrite -mulgA !(rH_Pmul x) ?rPmul //.
+ by rewrite -mulgA invMg -mulgA mulKg.
+have mu_Hmul x y z: x \in G -> y \in H -> mu x (y * z) = mu x z.
+ move=> Gx Hy; congr fmod; rewrite (mulgA x) (conjgCV x) -mulgA 2?rH_Hmul //.
+ by rewrite -mem_conjg (normP _) ?nHG.
+have{mu_Hmul} nu_Hmul y z: y \in H -> nu (y * z) = nu z.
+ move=> Hy; apply: eq_bigr => _ /rcosetsP[x Gx ->]; apply: mu_Hmul y z _ Hy.
+ by rewrite -(groupMl _ (subsetP sPG _ (PpP x))) mulgKV.
+have cocycle_mu: {in G & &, forall x y z,
+ mu (x * y)%g z + mu x y ^@ z = mu y z + mu x (y * z)%g}%R.
+- move=> x y z Gx Gy Gz; apply: val_inj.
+ apply: (mulgI (rH x * rH y * rH z)).
+ rewrite -(actrH _ _ Gz) addrC fmvalA fmvalJ ?nHG ?GrH //.
+ rewrite mulgA -(mulgA _ (rH z)) -conjgC mulgA -!rHmul ?groupM //.
+ by rewrite mulgA -mulgA -2!(mulgA (rH x)) -!rHmul ?groupM.
+move: mu => mu in rHmul mu_Pmul cocycle_mu nu nu_Hmul.
+have{cocycle_mu} cocycle_nu: {in G &, forall y z,
+ nu z + nu y ^@ z = mu y z *+ #|G : P| + nu (y * z)%g}%R.
+- move=> y z Gy Gz; rewrite /= (actr_sum z) /=.
+ have ->: (nu z = \sum_(Px in rcosets P G) mu (repr Px * y)%g z)%R.
+ rewrite /nu (reindex_acts _ (actsRs_rcosets P G) Gy) /=.
+ apply: eq_bigr => _ /rcosetsP[x Gx /= ->].
+ rewrite rcosetE -rcosetM.
+ case: repr_rcosetP=> p1 Pp1; case: repr_rcosetP=> p2 Pp2.
+ by rewrite -mulgA [x * y]lock !mu_Pmul.
+ rewrite -sumr_const -!big_split /=; apply: eq_bigr => _ /rcosetsP[x Gx ->].
+ rewrite -cocycle_mu //; case: repr_rcosetP => p1 Pp1.
+ by rewrite groupMr // (subsetP sPG).
+move: nu => nu in nu_Hmul cocycle_nu.
+pose f x := rH x * val (nu x *+ m)%R.
+have{cocycle_nu} fM: {in G &, {morph f : x y / x * y}}.
+ move=> x y Gx Gy; rewrite /f ?rHmul // -3!mulgA; congr (_ * _).
+ rewrite (mulgA _ (rH y)) (conjgC _ (rH y)) -mulgA; congr (_ * _).
+ rewrite -fmvalJ ?actrH ?nHG ?GrH // -!fmvalA actZr -mulrnDl.
+ rewrite -(addrC (nu y)) cocycle_nu // mulrnDl !fmvalA; congr (_ * _).
+ by rewrite !fmvalZ expgK ?fmodP.
+exists (Morphism fM @* G)%G; apply/complP; split.
+ apply/trivgP/subsetP=> x /setIP[Hx /morphimP[y _ Gy eq_x]].
+ apply/set1P; move: Hx; rewrite {x}eq_x /= groupMr ?subgP //.
+ rewrite -{1}(mulgKV y (rH y)) groupMl -?mem_rcoset // => Hy.
+ by rewrite -(mulg1 y) /f nu_Hmul // rH_Hmul //; exact: (morph1 (Morphism fM)).
+apply/setP=> x; apply/mulsgP/idP=> [[h y Hh fy ->{x}] | Gx].
+ rewrite groupMl; last exact: (subsetP sHG).
+ case/morphimP: fy => z _ Gz ->{h Hh y}.
+ by rewrite /= /f groupMl ?GrH // (subsetP sHG) ?fmodP.
+exists (x * (f x)^-1) (f x); last first; first by rewrite mulgKV.
+ by apply/morphimP; exists x.
+rewrite -groupV invMg invgK -mulgA (conjgC (val _)) mulgA.
+by rewrite groupMl -(mem_rcoset, mem_conjg) // (normP _) ?nHG ?fmodP.
+Qed.
+
+Theorem Gaschutz_transitive : {in [complements to H in G] &,
+ forall K L, K :&: P = L :&: P -> exists2 x, x \in H & L :=: K :^ x}.
+Proof.
+move=> K L /=; set Q := K :&: P => /complP[tiHK eqHK] cpHL QeqLP.
+have [trHL eqHL] := complP cpHL.
+pose nu x := fmod (divgr H L x^-1).
+have sKG: {subset K <= G} by apply/subsetP; rewrite -eqHK mulG_subr.
+have sLG: {subset L <= G} by apply/subsetP; rewrite -eqHL mulG_subr.
+have val_nu x: x \in G -> val (nu x) = divgr H L x^-1.
+ by move=> Gx; rewrite fmodK // mem_divgr // eqHL groupV.
+have nu_cocycle: {in G &, forall x y, nu (x * y)%g = nu x ^@ y + nu y}%R.
+ move=> x y Gx Gy; apply: val_inj; rewrite fmvalA fmvalJ ?nHG //.
+ rewrite !val_nu ?groupM // /divgr conjgE !mulgA mulgK.
+ by rewrite !(invMg, remgrM cpHL) ?groupV ?mulgA.
+have nuL x: x \in L -> nu x = 0%R.
+ move=> Lx; apply: val_inj; rewrite val_nu ?sLG //.
+ by rewrite /divgr remgr_id ?groupV ?mulgV.
+exists (fmval ((\sum_(X in rcosets Q K) nu (repr X)) *+ m)).
+ exact: fmodP.
+apply/eqP; rewrite eq_sym eqEcard; apply/andP; split; last first.
+ by rewrite cardJg -(leq_pmul2l (cardG_gt0 H)) -!TI_cardMg // eqHL eqHK.
+apply/subsetP=> _ /imsetP[x Kx ->]; rewrite conjgE mulgA (conjgC _ x).
+have Gx: x \in G by rewrite sKG.
+rewrite conjVg -mulgA -fmvalJ ?nHG // -fmvalN -fmvalA (_ : _ + _ = nu x)%R.
+ by rewrite val_nu // mulKVg groupV mem_remgr // eqHL groupV.
+rewrite actZr -!mulNrn -mulrnDl actr_sum.
+rewrite addrC (reindex_acts _ (actsRs_rcosets _ K) Kx) -sumrB /= -/Q.
+rewrite (eq_bigr (fun _ => nu x)) => [|_ /imsetP[y Ky ->]]; last first.
+ rewrite !rcosetE -rcosetM QeqLP.
+ case: repr_rcosetP => z /setIP[Lz _]; case: repr_rcosetP => t /setIP[Lt _].
+ rewrite !nu_cocycle ?groupM ?(sKG y) // ?sLG //.
+ by rewrite (nuL z) ?(nuL t) // !act0r !add0r addrC addKr.
+apply: val_inj; rewrite sumr_const !fmvalZ.
+rewrite -{2}(expgK coHiPG (fmodP (nu x))); congr (_ ^+ _ ^+ _).
+rewrite -[#|_|]divgS ?subsetIl // -(divnMl (cardG_gt0 H)).
+rewrite -!TI_cardMg //; last by rewrite setIA setIAC (setIidPl sHP).
+by rewrite group_modl // eqHK (setIidPr sPG) divgS.
+Qed.
+
+End Gaschutz.
+
+(* This is the TI part of B & G, Proposition 1.6(d). *)
+(* We go with B & G rather than Aschbacher and will derive 1.6(e) from (d), *)
+(* rather than the converse, because the derivation of 24.6 from 24.3 in *)
+(* Aschbacher requires a separate reduction to p-groups to yield 1.6(d), *)
+(* making it altogether longer than the direct Gaschutz-style proof. *)
+(* This Lemma is used in maximal.v for the proof of Aschbacher 24.7. *)
+Lemma coprime_abel_cent_TI (gT : finGroupType) (A G : {group gT}) :
+ A \subset 'N(G) -> coprime #|G| #|A| -> abelian G -> 'C_[~: G, A](A) = 1.
+Proof.
+move=> nGA coGA abG; pose f x := val (\sum_(a in A) fmod abG x ^@ a)%R.
+have fM: {in G &, {morph f : x y / x * y}}.
+ move=> x y Gx Gy /=; rewrite -fmvalA -big_split /=; congr (fmval _).
+ by apply: eq_bigr => a Aa; rewrite fmodM // actAr.
+have nfA x a: a \in A -> f (x ^ a) = f x.
+ move=> Aa; rewrite {2}/f (reindex_inj (mulgI a)) /=; congr (fmval _).
+ apply: eq_big => [b | b Ab]; first by rewrite groupMl.
+ by rewrite -!fmodJ ?groupM ?(subsetP nGA) // conjgM.
+have kerR: [~: G, A] \subset 'ker (Morphism fM).
+ rewrite gen_subG; apply/subsetP=> xa; case/imset2P=> x a Gx Aa -> {xa}.
+ have Gxa: x ^ a \in G by rewrite memJ_norm ?(subsetP nGA).
+ rewrite commgEl; apply/kerP; rewrite (groupM, morphM) ?(groupV, morphV) //=.
+ by rewrite nfA ?mulVg.
+apply/trivgP; apply/subsetP=> x /setIP[Rx cAx]; apply/set1P.
+have Gx: x \in G by apply: subsetP Rx; rewrite commg_subl.
+rewrite -(expgK coGA Gx) (_ : x ^+ _ = 1) ?expg1n //.
+rewrite -(fmodK abG Gx) -fmvalZ -(mker (subsetP kerR x Rx)); congr fmval.
+rewrite -GRing.sumr_const; apply: eq_bigr => a Aa.
+by rewrite -fmodJ ?(subsetP nGA) // /conjg (centP cAx) // mulKg.
+Qed.
+
+Section Transfer.
+
+Variables (gT aT : finGroupType) (G H : {group gT}).
+Variable alpha : {morphism H >-> aT}.
+
+Hypotheses (sHG : H \subset G) (abelA : abelian (alpha @* H)).
+
+Local Notation HG := (rcosets (gval H) (gval G)).
+
+Fact transfer_morph_subproof : H \subset alpha @*^-1 (alpha @* H).
+Proof. by rewrite -sub_morphim_pre. Qed.
+
+Let fmalpha := restrm transfer_morph_subproof (fmod abelA \o alpha).
+
+Let V (rX : {set gT} -> gT) g :=
+ \sum_(Hx in rcosets H G) fmalpha (rX Hx * g * (rX (Hx :* g))^-1).
+
+Definition transfer g := V repr g.
+
+(* This is Aschbacher (37.2). *)
+Lemma transferM : {in G &, {morph transfer : x y / (x * y)%g >-> x + y}}.
+Proof.
+move=> s t Gs Gt /=.
+rewrite [transfer t](reindex_acts 'Rs _ Gs) ?actsRs_rcosets //= -big_split /=.
+apply: eq_bigr => _ /rcosetsP[x Gx ->]; rewrite !rcosetE -!rcosetM.
+rewrite -zmodMgE -morphM -?mem_rcoset; first by rewrite !mulgA mulgKV rcosetM.
+ by rewrite rcoset_repr rcosetM mem_rcoset mulgK mem_repr_rcoset.
+by rewrite rcoset_repr (rcosetM _ _ t) mem_rcoset mulgK mem_repr_rcoset.
+Qed.
+
+Canonical transfer_morphism := Morphism transferM.
+
+(* This is Aschbacher (37.1). *)
+Lemma transfer_indep X (rX := transversal_repr 1 X) :
+ is_transversal X HG G -> {in G, transfer =1 V rX}.
+Proof.
+move=> trX g Gg; have mem_rX := repr_mem_pblock trX 1; rewrite -/rX in mem_rX.
+apply: (addrI (\sum_(Hx in HG) fmalpha (repr Hx * (rX Hx)^-1))).
+rewrite {1}(reindex_acts 'Rs _ Gg) ?actsRs_rcosets // -!big_split /=.
+apply: eq_bigr => _ /rcosetsP[x Gx ->]; rewrite !rcosetE -!rcosetM.
+case: repr_rcosetP => h1 Hh1; case: repr_rcosetP => h2 Hh2.
+have: H :* (x * g) \in rcosets H G by rewrite -rcosetE mem_imset ?groupM.
+have: H :* x \in rcosets H G by rewrite -rcosetE mem_imset.
+case/mem_rX/rcosetP=> h3 Hh3 -> /mem_rX/rcosetP[h4 Hh4 ->].
+rewrite -!(mulgA h1) -!(mulgA h2) -!(mulgA h3) !(mulKVg, invMg).
+by rewrite addrC -!zmodMgE -!morphM ?groupM ?groupV // -!mulgA !mulKg.
+Qed.
+
+Section FactorTransfer.
+
+Variable g : gT.
+Hypothesis Gg : g \in G.
+
+Let sgG : <[g]> \subset G. Proof. by rewrite cycle_subG. Qed.
+Let H_g_rcosets x : {set {set gT}} := rcosets (H :* x) <[g]>.
+Let n_ x := #|<[g]> : H :* x|.
+
+Lemma mulg_exp_card_rcosets x : x * (g ^+ n_ x) \in H :* x.
+Proof.
+rewrite /n_ /indexg -orbitRs -pcycle_actperm ?inE //.
+rewrite -{2}(iter_pcycle (actperm 'Rs g) (H :* x)) -permX -morphX ?inE //.
+by rewrite actpermE //= rcosetE -rcosetM rcoset_refl.
+Qed.
+
+Let HGg : {set {set {set gT}}} := orbit 'Rs <[g]> @: HG.
+
+Let partHG : partition HG G := rcosets_partition sHG.
+Let actsgHG : [acts <[g]>, on HG | 'Rs].
+Proof. exact: subset_trans sgG (actsRs_rcosets H G). Qed.
+Let partHGg : partition HGg HG := orbit_partition actsgHG.
+
+Let injHGg : {in HGg &, injective cover}.
+Proof. by have [] := partition_partition partHG partHGg. Qed.
+
+Let defHGg : HG :* <[g]> = cover @: HGg.
+Proof.
+rewrite -imset_comp [_ :* _]imset2_set1r; apply: eq_imset => Hx /=.
+by rewrite cover_imset -curry_imset2r.
+Qed.
+
+Lemma rcosets_cycle_partition : partition (HG :* <[g]>) G.
+Proof. by rewrite defHGg; have [] := partition_partition partHG partHGg. Qed.
+
+Variable X : {set gT}.
+Hypothesis trX : is_transversal X (HG :* <[g]>) G.
+
+Let sXG : {subset X <= G}. Proof. exact/subsetP/(transversal_sub trX). Qed.
+
+Lemma rcosets_cycle_transversal : H_g_rcosets @: X = HGg.
+Proof.
+have sHXgHGg x: x \in X -> H_g_rcosets x \in HGg.
+ by move/sXG=> Gx; apply: mem_imset; rewrite -rcosetE mem_imset.
+apply/setP=> Hxg; apply/imsetP/idP=> [[x /sHXgHGg HGgHxg -> //] | HGgHxg].
+have [_ /rcosetsP[z Gz ->] ->] := imsetP HGgHxg.
+pose Hzg := H :* z * <[g]>; pose x := transversal_repr 1 X Hzg.
+have HGgHzg: Hzg \in HG :* <[g]>.
+ by rewrite mem_mulg ?set11 // -rcosetE mem_imset.
+have Hzg_x: x \in Hzg by rewrite (repr_mem_pblock trX).
+exists x; first by rewrite (repr_mem_transversal trX).
+case/mulsgP: Hzg_x => y u /rcoset_transl <- /(orbit_act 'Rs) <- -> /=.
+by rewrite rcosetE -rcosetM.
+Qed.
+
+Local Notation defHgX := rcosets_cycle_transversal.
+
+Let injHg: {in X &, injective H_g_rcosets}.
+Proof.
+apply/imset_injP; rewrite defHgX (card_transversal trX) defHGg.
+by rewrite (card_in_imset injHGg).
+Qed.
+
+Lemma sum_index_rcosets_cycle : (\sum_(x in X) n_ x)%N = #|G : H|.
+Proof. by rewrite [#|G : H|](card_partition partHGg) -defHgX big_imset. Qed.
+
+Lemma transfer_cycle_expansion :
+ transfer g = \sum_(x in X) fmalpha ((g ^+ n_ x) ^ x^-1).
+Proof.
+pose Y := \bigcup_(x in X) [set x * g ^+ i | i : 'I_(n_ x)].
+pose rY := transversal_repr 1 Y.
+pose pcyc x := pcycle (actperm 'Rs g) (H :* x).
+pose traj x := traject (actperm 'Rs g) (H :* x) #|pcyc x|.
+have Hgr_eq x: H_g_rcosets x = pcyc x.
+ by rewrite /H_g_rcosets -orbitRs -pcycle_actperm ?inE.
+have pcyc_eq x: pcyc x =i traj x by exact: pcycle_traject.
+have uniq_traj x: uniq (traj x) by apply: uniq_traject_pcycle.
+have n_eq x: n_ x = #|pcyc x| by rewrite -Hgr_eq.
+have size_traj x: size (traj x) = n_ x by rewrite n_eq size_traject.
+have nth_traj x j: j < n_ x -> nth (H :* x) (traj x) j = H :* (x * g ^+ j).
+ move=> lt_j_x; rewrite nth_traject -?n_eq //.
+ by rewrite -permX -morphX ?inE // actpermE //= rcosetE rcosetM.
+have sYG: Y \subset G.
+ apply/bigcupsP=> x Xx; apply/subsetP=> _ /imsetP[i _ ->].
+ by rewrite groupM ?groupX // sXG.
+have trY: is_transversal Y HG G.
+ apply/and3P; split=> //; apply/forall_inP=> Hy.
+ have /and3P[/eqP <- _ _] := partHGg; rewrite -defHgX cover_imset.
+ case/bigcupP=> x Xx; rewrite Hgr_eq pcyc_eq => /trajectP[i].
+ rewrite -n_eq -permX -morphX ?in_setT // actpermE /= rcosetE -rcosetM => lti.
+ set y := x * _ => ->{Hy}; pose oi := Ordinal lti.
+ have Yy: y \in Y by apply/bigcupP; exists x => //; apply/imsetP; exists oi.
+ apply/cards1P; exists y; apply/esym/eqP.
+ rewrite eqEsubset sub1set inE Yy rcoset_refl.
+ apply/subsetP=> _ /setIP[/bigcupP[x' Xx' /imsetP[j _ ->]] Hy_x'gj].
+ have eq_xx': x = x'.
+ apply: (pblock_inj trX) => //; have /andP[/and3P[_ tiX _] _] := trX.
+ have HGgHyg: H :* y * <[g]> \in HG :* <[g]>.
+ by rewrite mem_mulg ?set11 // -rcosetE mem_imset ?(subsetP sYG).
+ rewrite !(def_pblock tiX HGgHyg) //.
+ by rewrite -[x'](mulgK (g ^+ j)) mem_mulg // groupV mem_cycle.
+ by rewrite -[x](mulgK (g ^+ i)) mem_mulg ?rcoset_refl // groupV mem_cycle.
+ apply/set1P; rewrite /y eq_xx'; congr (_ * _ ^+ _) => //; apply/eqP.
+ rewrite -(@nth_uniq _ (H :* x) (traj x)) ?size_traj // ?eq_xx' //.
+ by rewrite !nth_traj ?(rcoset_transl Hy_x'gj) // -eq_xx'.
+have rYE x i : x \in X -> i < n_ x -> rY (H :* x :* g ^+ i) = x * g ^+ i.
+ move=> Xx lt_i_x; rewrite -rcosetM; apply: (canLR_in (pblockK trY 1)).
+ by apply/bigcupP; exists x => //; apply/imsetP; exists (Ordinal lt_i_x).
+ apply/esym/def_pblock; last exact: rcoset_refl; first by case/and3P: partHG.
+ by rewrite -rcosetE mem_imset ?groupM ?groupX // sXG.
+rewrite (transfer_indep trY Gg) /V -/rY (set_partition_big _ partHGg) /=.
+rewrite -defHgX big_imset /=; last first.
+ apply/imset_injP; rewrite defHgX (card_transversal trX) defHGg.
+ by rewrite (card_in_imset injHGg).
+apply eq_bigr=> x Xx; rewrite Hgr_eq (eq_bigl _ _ (pcyc_eq x)) -big_uniq //=.
+have n_gt0: 0 < n_ x by rewrite indexg_gt0.
+rewrite /traj -n_eq; case def_n: (n_ x) (n_gt0) => // [n] _.
+rewrite conjgE invgK -{1}[H :* x]rcoset1 -{1}(expg0 g).
+elim: {1 3}n 0%N (addn0 n) => [|m IHm] i def_i /=.
+ rewrite big_seq1 {i}[i]def_i rYE // ?def_n //.
+ rewrite -(mulgA _ _ g) -rcosetM -expgSr -[(H :* x) :* _]rcosetE.
+ rewrite -actpermE morphX ?inE // permX // -{2}def_n n_eq iter_pcycle mulgA.
+ by rewrite -[H :* x]rcoset1 (rYE _ 0%N) ?mulg1.
+rewrite big_cons rYE //; last by rewrite def_n -def_i ltnS leq_addl.
+rewrite permE /= rcosetE -rcosetM -(mulgA _ _ g) -expgSr.
+rewrite addSnnS in def_i; rewrite IHm //.
+rewrite rYE //; last by rewrite def_n -def_i ltnS leq_addl.
+by rewrite mulgV [fmalpha 1]morph1 add0r.
+Qed.
+
+End FactorTransfer.
+
+End Transfer.
diff --git a/mathcomp/solvable/frobenius.v b/mathcomp/solvable/frobenius.v
new file mode 100644
index 0000000..492c802
--- /dev/null
+++ b/mathcomp/solvable/frobenius.v
@@ -0,0 +1,794 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat div fintype bigop prime.
+Require Import finset fingroup morphism perm action quotient gproduct.
+Require Import cyclic center pgroup nilpotent sylow hall abelian.
+
+(******************************************************************************)
+(* Definition of Frobenius groups, some basic results, and the Frobenius *)
+(* theorem on the number of solutions of x ^+ n = 1. *)
+(* semiregular K H <-> *)
+(* the internal action of H on K is semiregular, i.e., no nontrivial *)
+(* elements of H and K commute; note that this is actually a symmetric *)
+(* condition. *)
+(* semiprime K H <-> *)
+(* the internal action of H on K is "prime", i.e., an element of K that *)
+(* centralises a nontrivial element of H must actually centralise all *)
+(* of H. *)
+(* normedTI A G L <=> *)
+(* A is nonempty, strictly disjoint from its conjugates in G, and has *)
+(* normaliser L in G. *)
+(* [Frobenius G = K ><| H] <=> *)
+(* G is (isomorphic to) a Frobenius group with kernel K and complement *)
+(* H. This is an effective predicate (in bool), which tests the *)
+(* equality with the semidirect product, and then the fact that H is a *)
+(* proper self-normalizing TI-subgroup of G. *)
+(* [Frobenius G with kernel H] <=> *)
+(* G is (isomorphic to) a Frobenius group with kernel K; same as above, *)
+(* but without the semi-direct product. *)
+(* [Frobenius G with complement H] <=> *)
+(* G is (isomorphic to) a Frobenius group with complement H; same as *)
+(* above, but without the semi-direct product. The proof that this form *)
+(* is equivalent to the above (i.e., the existence of Frobenius *)
+(* kernels) requires chareacter theory and will only be proved in the *)
+(* vcharacter module. *)
+(* [Frobenius G] <=> G is a Frobenius group. *)
+(* Frobenius_action G H S to <-> *)
+(* The action to of G on S defines an isomorphism of G with a *)
+(* (permutation) Frobenius group, i.e., to is faithful and transitive *)
+(* on S, no nontrivial element of G fixes more than one point in S, and *)
+(* H is the stabilizer of some element of S, and non-trivial. Thus, *)
+(* Frobenius_action G H S 'P *)
+(* asserts that G is a Frobenius group in the classic sense. *)
+(* has_Frobenius_action G H <-> *)
+(* Frobenius_action G H S to holds for some sT : finType, S : {set st} *)
+(* and to : {action gT &-> sT}. This is a predicate in Prop, but is *)
+(* exactly reflected by [Frobenius G with complement H] : bool. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Definitions.
+
+Variable gT : finGroupType.
+Implicit Types A G K H L : {set gT}.
+
+(* Corresponds to "H acts on K in a regular manner" in B & G. *)
+Definition semiregular K H := {in H^#, forall x, 'C_K[x] = 1}.
+
+(* Corresponds to "H acts on K in a prime manner" in B & G. *)
+Definition semiprime K H := {in H^#, forall x, 'C_K[x] = 'C_K(H)}.
+
+Definition normedTI A G L := [&& A != set0, trivIset (A :^: G) & 'N_G(A) == L].
+
+Definition Frobenius_group_with_complement G H := (H != G) && normedTI H^# G H.
+
+Definition Frobenius_group G :=
+ [exists H : {group gT}, Frobenius_group_with_complement G H].
+
+Definition Frobenius_group_with_kernel_and_complement G K H :=
+ (K ><| H == G) && Frobenius_group_with_complement G H.
+
+Definition Frobenius_group_with_kernel G K :=
+ [exists H : {group gT}, Frobenius_group_with_kernel_and_complement G K H].
+
+Section FrobeniusAction.
+
+Variables G H : {set gT}.
+Variables (sT : finType) (S : {set sT}) (to : {action gT &-> sT}).
+
+Definition Frobenius_action :=
+ [/\ [faithful G, on S | to],
+ [transitive G, on S | to],
+ {in G^#, forall x, #|'Fix_(S | to)[x]| <= 1},
+ H != 1
+ & exists2 u, u \in S & H = 'C_G[u | to]].
+
+End FrobeniusAction.
+
+CoInductive has_Frobenius_action G H : Prop :=
+ HasFrobeniusAction sT S to of @Frobenius_action G H sT S to.
+
+End Definitions.
+
+Arguments Scope semiregular [_ group_scope group_scope].
+Arguments Scope semiprime [_ group_scope group_scope].
+Arguments Scope normedTI [_ group_scope group_scope group_scope].
+Arguments Scope Frobenius_group_with_complement [_ group_scope group_scope].
+Arguments Scope Frobenius_group [_ group_scope].
+Arguments Scope Frobenius_group_with_kernel [_ group_scope group_scope].
+Arguments Scope Frobenius_group_with_kernel_and_complement
+ [_ group_scope group_scope group_scope].
+Arguments Scope Frobenius_action
+ [_ group_scope group_scope _ group_scope action_scope].
+Arguments Scope has_Frobenius_action [_ group_scope group_scope].
+
+Notation "[ 'Frobenius' G 'with' 'complement' H ]" :=
+ (Frobenius_group_with_complement G H)
+ (at level 0, G at level 50, H at level 35,
+ format "[ 'Frobenius' G 'with' 'complement' H ]") : group_scope.
+
+Notation "[ 'Frobenius' G 'with' 'kernel' K ]" :=
+ (Frobenius_group_with_kernel G K)
+ (at level 0, G at level 50, K at level 35,
+ format "[ 'Frobenius' G 'with' 'kernel' K ]") : group_scope.
+
+Notation "[ 'Frobenius' G ]" :=
+ (Frobenius_group G)
+ (at level 0, G at level 50,
+ format "[ 'Frobenius' G ]") : group_scope.
+
+Notation "[ 'Frobenius' G = K ><| H ]" :=
+ (Frobenius_group_with_kernel_and_complement G K H)
+ (at level 0, G at level 50, K, H at level 35,
+ format "[ 'Frobenius' G = K ><| H ]") : group_scope.
+
+Section FrobeniusBasics.
+
+Variable gT : finGroupType.
+Implicit Types (A B : {set gT}) (G H K L R X : {group gT}).
+
+Lemma semiregular1l H : semiregular 1 H.
+Proof. by move=> x _ /=; rewrite setI1g. Qed.
+
+Lemma semiregular1r K : semiregular K 1.
+Proof. by move=> x; rewrite setDv inE. Qed.
+
+Lemma semiregular_sym H K : semiregular K H -> semiregular H K.
+Proof.
+move=> regH x /setD1P[ntx Kx]; apply: contraNeq ntx.
+rewrite -subG1 -setD_eq0 -setIDAC => /set0Pn[y /setIP[Hy cxy]].
+by rewrite (sameP eqP set1gP) -(regH y Hy) inE Kx cent1C.
+Qed.
+
+Lemma semiregularS K1 K2 A1 A2 :
+ K1 \subset K2 -> A1 \subset A2 -> semiregular K2 A2 -> semiregular K1 A1.
+Proof.
+move=> sK12 sA12 regKA2 x /setD1P[ntx /(subsetP sA12)A2x].
+by apply/trivgP; rewrite -(regKA2 x) ?inE ?ntx ?setSI.
+Qed.
+
+Lemma semiregular_prime H K : semiregular K H -> semiprime K H.
+Proof.
+move=> regH x Hx; apply/eqP; rewrite eqEsubset {1}regH // sub1G.
+by rewrite -cent_set1 setIS ?centS // sub1set; case/setD1P: Hx.
+Qed.
+
+Lemma semiprime_regular H K : semiprime K H -> 'C_K(H) = 1 -> semiregular K H.
+Proof. by move=> prKH tiKcH x Hx; rewrite prKH. Qed.
+
+Lemma semiprimeS K1 K2 A1 A2 :
+ K1 \subset K2 -> A1 \subset A2 -> semiprime K2 A2 -> semiprime K1 A1.
+Proof.
+move=> sK12 sA12 prKA2 x /setD1P[ntx A1x].
+apply/eqP; rewrite eqEsubset andbC -{1}cent_set1 setIS ?centS ?sub1set //=.
+rewrite -(setIidPl sK12) -!setIA prKA2 ?setIS ?centS //.
+by rewrite !inE ntx (subsetP sA12).
+Qed.
+
+Lemma cent_semiprime H K X :
+ semiprime K H -> X \subset H -> X :!=: 1 -> 'C_K(X) = 'C_K(H).
+Proof.
+move=> prKH sXH /trivgPn[x Xx ntx]; apply/eqP.
+rewrite eqEsubset -{1}(prKH x) ?inE ?(subsetP sXH) ?ntx //=.
+by rewrite -cent_cycle !setIS ?centS ?cycle_subG.
+Qed.
+
+Lemma stab_semiprime H K X :
+ semiprime K H -> X \subset K -> 'C_H(X) != 1 -> 'C_H(X) = H.
+Proof.
+move=> prKH sXK ntCHX; apply/setIidPl; rewrite centsC -subsetIidl.
+rewrite -{2}(setIidPl sXK) -setIA -(cent_semiprime prKH _ ntCHX) ?subsetIl //.
+by rewrite !subsetI subxx sXK centsC subsetIr.
+Qed.
+
+Lemma cent_semiregular H K X :
+ semiregular K H -> X \subset H -> X :!=: 1 -> 'C_K(X) = 1.
+Proof.
+move=> regKH sXH /trivgPn[x Xx ntx]; apply/trivgP.
+rewrite -(regKH x) ?inE ?(subsetP sXH) ?ntx ?setIS //=.
+by rewrite -cent_cycle centS ?cycle_subG.
+Qed.
+
+Lemma regular_norm_dvd_pred K H :
+ H \subset 'N(K) -> semiregular K H -> #|H| %| #|K|.-1.
+Proof.
+move=> nKH regH; have actsH: [acts H, on K^# | 'J] by rewrite astabsJ normD1.
+rewrite (cardsD1 1 K) group1 -(acts_sum_card_orbit actsH) /=.
+rewrite (eq_bigr (fun _ => #|H|)) ?sum_nat_const ?dvdn_mull //.
+move=> _ /imsetP[x /setIdP[ntx Kx] ->]; rewrite card_orbit astab1J.
+rewrite ['C_H[x]](trivgP _) ?indexg1 //=.
+apply/subsetP=> y /setIP[Hy cxy]; apply: contraR ntx => nty.
+by rewrite -[[set 1]](regH y) inE ?nty // Kx cent1C.
+
+Qed.
+
+Lemma regular_norm_coprime K H :
+ H \subset 'N(K) -> semiregular K H -> coprime #|K| #|H|.
+Proof.
+move=> nKH regH.
+by rewrite (coprime_dvdr (regular_norm_dvd_pred nKH regH)) ?coprimenP.
+Qed.
+
+Lemma semiregularJ K H x : semiregular K H -> semiregular (K :^ x) (H :^ x).
+Proof.
+move=> regH yx; rewrite -conjD1g => /imsetP[y Hy ->].
+by rewrite cent1J -conjIg regH ?conjs1g.
+Qed.
+
+Lemma semiprimeJ K H x : semiprime K H -> semiprime (K :^ x) (H :^ x).
+Proof.
+move=> prH yx; rewrite -conjD1g => /imsetP[y Hy ->].
+by rewrite cent1J centJ -!conjIg prH.
+Qed.
+
+Lemma normedTI_P A G L :
+ reflect [/\ A != set0, L \subset 'N_G(A)
+ & {in G, forall g, ~~ [disjoint A & A :^ g] -> g \in L}]
+ (normedTI A G L).
+Proof.
+apply: (iffP and3P) => [[nzA /trivIsetP tiAG /eqP <-] | [nzA sLN tiAG]].
+ split=> // g Gg; rewrite inE Gg (sameP normP eqP) /= eq_sym; apply: contraR.
+ by apply: tiAG; rewrite ?mem_orbit ?orbit_refl.
+have [/set0Pn[a Aa] /subsetIP[_ nAL]] := (nzA, sLN); split=> //; last first.
+ rewrite eqEsubset sLN andbT; apply/subsetP=> x /setIP[Gx nAx].
+ by apply/tiAG/pred0Pn=> //; exists a; rewrite /= (normP nAx) Aa.
+apply/trivIsetP=> _ _ /imsetP[x Gx ->] /imsetP[y Gy ->]; apply: contraR.
+rewrite -setI_eq0 -(mulgKV x y) conjsgM; set g := (y * x^-1)%g.
+have Gg: g \in G by rewrite groupMl ?groupV.
+rewrite -conjIg (inj_eq (act_inj 'Js x)) (eq_sym A) (sameP eqP normP).
+by rewrite -cards_eq0 cardJg cards_eq0 setI_eq0 => /tiAG/(subsetP nAL)->.
+Qed.
+Implicit Arguments normedTI_P [A G L].
+
+Lemma normedTI_memJ_P A G L :
+ reflect [/\ A != set0, L \subset G
+ & {in A & G, forall a g, (a ^ g \in A) = (g \in L)}]
+ (normedTI A G L).
+Proof.
+apply: (iffP normedTI_P) => [[-> /subsetIP[sLG nAL] tiAG] | [-> sLG tiAG]].
+ split=> // a g Aa Gg; apply/idP/idP=> [Aag | Lg]; last first.
+ by rewrite memJ_norm ?(subsetP nAL).
+ by apply/tiAG/pred0Pn=> //; exists (a ^ g)%g; rewrite /= Aag memJ_conjg.
+split=> // [ | g Gg /pred0Pn[ag /=]]; last first.
+ by rewrite andbC => /andP[/imsetP[a Aa ->]]; rewrite tiAG.
+apply/subsetP=> g Lg; have Gg := subsetP sLG g Lg.
+by rewrite !inE Gg; apply/subsetP=> _ /imsetP[a Aa ->]; rewrite tiAG.
+Qed.
+
+Lemma partition_class_support A G :
+ A != set0 -> trivIset (A :^: G) -> partition (A :^: G) (class_support A G).
+Proof.
+rewrite /partition cover_imset -class_supportEr eqxx => nzA ->.
+by apply: contra nzA => /imsetP[x _ /eqP]; rewrite eq_sym -!cards_eq0 cardJg.
+Qed.
+
+Lemma partition_normedTI A G L :
+ normedTI A G L -> partition (A :^: G) (class_support A G).
+Proof. by case/and3P=> ntA tiAG _; apply: partition_class_support. Qed.
+
+Lemma card_support_normedTI A G L :
+ normedTI A G L -> #|class_support A G| = (#|A| * #|G : L|)%N.
+Proof.
+case/and3P=> ntA tiAG /eqP <-; rewrite -card_conjugates mulnC.
+apply: card_uniform_partition (partition_class_support ntA tiAG).
+by move=> _ /imsetP[y _ ->]; rewrite cardJg.
+Qed.
+
+Lemma normedTI_S A B G L :
+ A != set0 -> L \subset 'N(A) -> A \subset B -> normedTI B G L ->
+ normedTI A G L.
+Proof.
+move=> nzA /subsetP nAL /subsetP sAB /normedTI_memJ_P[nzB sLG tiB].
+apply/normedTI_memJ_P; split=> // a x Aa Gx.
+by apply/idP/idP => [Aax | /nAL/memJ_norm-> //]; rewrite -(tiB a) ?sAB.
+Qed.
+
+Lemma cent1_normedTI A G L :
+ normedTI A G L -> {in A, forall x, 'C_G[x] \subset L}.
+Proof.
+case/normedTI_memJ_P=> [_ _ tiAG] x Ax; apply/subsetP=> y /setIP[Gy cxy].
+by rewrite -(tiAG x) // /(x ^ y) -(cent1P cxy) mulKg.
+Qed.
+
+Lemma Frobenius_actionP G H :
+ reflect (has_Frobenius_action G H) [Frobenius G with complement H].
+Proof.
+apply: (iffP andP) => [[neqHG] | [sT S to [ffulG transG regG ntH [u Su defH]]]].
+ case/normedTI_P=> nzH /subsetIP[sHG _] tiHG.
+ suffices: Frobenius_action G H (rcosets H G) 'Rs by exact: HasFrobeniusAction.
+ pose Hfix x := 'Fix_(rcosets H G | 'Rs)[x].
+ have regG: {in G^#, forall x, #|Hfix x| <= 1}.
+ move=> x /setD1P[ntx Gx].
+ apply: wlog_neg; rewrite -ltnNge => /ltnW/card_gt0P/=[Hy].
+ rewrite -(cards1 Hy) => /setIP[/imsetP[y Gy ->{Hy}] cHyx].
+ apply/subset_leq_card/subsetP=> _ /setIP[/imsetP[z Gz ->] cHzx].
+ rewrite -!sub_astab1 !astab1_act !sub1set astab1Rs in cHyx cHzx *.
+ rewrite !rcosetE; apply/set1P/rcoset_transl; rewrite mem_rcoset.
+ apply: tiHG; [by rewrite !in_group | apply/pred0Pn; exists (x ^ y^-1)].
+ by rewrite conjD1g !inE conjg_eq1 ntx -mem_conjg cHyx conjsgM memJ_conjg.
+ have ntH: H :!=: 1 by rewrite -subG1 -setD_eq0.
+ split=> //; first 1 last; first exact: transRs_rcosets.
+ by exists (H : {set gT}); rewrite ?orbit_refl // astab1Rs (setIidPr sHG).
+ apply/subsetP=> y /setIP[Gy cHy]; apply: contraR neqHG => nt_y.
+ rewrite (index1g sHG) //; apply/eqP; rewrite eqn_leq indexg_gt0 andbT.
+ apply: leq_trans (regG y _); last by rewrite setDE 2!inE Gy nt_y /=.
+ by rewrite /Hfix (setIidPl _) -1?astabC ?sub1set.
+have sHG: H \subset G by rewrite defH subsetIl.
+split.
+ apply: contraNneq ntH => /= defG.
+ suffices defS: S = [set u] by rewrite -(trivgP ffulG) /= defS defH.
+ apply/eqP; rewrite eq_sym eqEcard sub1set Su.
+ by rewrite -(atransP transG u Su) card_orbit -defH defG indexgg cards1.
+apply/normedTI_P; rewrite setD_eq0 subG1 normD1 subsetI sHG normG.
+split=> // x Gx; rewrite -setI_eq0 conjD1g defH inE Gx conjIg conjGid //.
+rewrite -setDIl -setIIr -astab1_act setDIl => /set0Pn[y /setIP[Gy /setD1P[_]]].
+case/setIP; rewrite 2!(sameP astab1P afix1P) => cuy cuxy; apply/astab1P.
+apply: contraTeq (regG y Gy) => cu'x.
+rewrite (cardD1 u) (cardD1 (to u x)) inE Su cuy inE /= inE cu'x cuxy.
+by rewrite (actsP (atrans_acts transG)) ?Su.
+Qed.
+
+Section FrobeniusProperties.
+
+Variables G H K : {group gT}.
+Hypothesis frobG : [Frobenius G = K ><| H].
+
+Lemma FrobeniusWker : [Frobenius G with kernel K].
+Proof. by apply/existsP; exists H. Qed.
+
+Lemma FrobeniusWcompl : [Frobenius G with complement H].
+Proof. by case/andP: frobG. Qed.
+
+Lemma FrobeniusW : [Frobenius G].
+Proof. by apply/existsP; exists H; exact: FrobeniusWcompl. Qed.
+
+Lemma Frobenius_context :
+ [/\ K ><| H = G, K :!=: 1, H :!=: 1, K \proper G & H \proper G].
+Proof.
+have [/eqP defG neqHG ntH _] := and4P frobG; rewrite setD_eq0 subG1 in ntH.
+have ntK: K :!=: 1 by apply: contraNneq neqHG => K1; rewrite -defG K1 sdprod1g.
+rewrite properEcard properEneq neqHG; have /mulG_sub[-> ->] := sdprodW defG.
+by rewrite -(sdprod_card defG) ltn_Pmulr ?cardG_gt1.
+Qed.
+
+Lemma Frobenius_partition : partition (gval K |: (H^# :^: K)) G.
+Proof.
+have [/eqP defG _ tiHG] := and3P frobG; have [_ tiH1G /eqP defN] := and3P tiHG.
+have [[_ /mulG_sub[sKG sHG] nKH tiKH] mulHK] := (sdprodP defG, sdprodWC defG).
+set HG := H^# :^: K; set KHG := _ |: _.
+have defHG: HG = H^# :^: G.
+ have: 'C_G[H^# | 'Js] * K = G by rewrite astab1Js defN mulHK.
+ move/subgroup_transitiveP/atransP.
+ by apply; rewrite ?atrans_orbit ?orbit_refl.
+have /and3P[defHK _ nzHG] := partition_normedTI tiHG.
+rewrite -defHG in defHK nzHG tiH1G.
+have [tiKHG HG'K]: trivIset KHG /\ gval K \notin HG.
+ apply: trivIsetU1 => // _ /imsetP[x Kx ->]; rewrite -setI_eq0.
+ by rewrite -(conjGid Kx) -conjIg setIDA tiKH setDv conj0g.
+rewrite /partition andbC tiKHG !inE negb_or nzHG eq_sym -card_gt0 cardG_gt0 /=.
+rewrite eqEcard; apply/andP; split.
+ rewrite /cover big_setU1 //= subUset sKG -/(cover HG) (eqP defHK).
+ by rewrite class_support_subG // (subset_trans _ sHG) ?subD1set.
+rewrite -(eqnP tiKHG) big_setU1 //= (eqnP tiH1G) (eqP defHK).
+rewrite (card_support_normedTI tiHG) -(Lagrange sHG) (cardsD1 1) group1 mulSn.
+by rewrite leq_add2r -mulHK indexMg -indexgI tiKH indexg1.
+Qed.
+
+Lemma Frobenius_cent1_ker : {in K^#, forall x, 'C_G[x] \subset K}.
+Proof.
+have [/eqP defG _ /normedTI_memJ_P[_ _ tiHG]] := and3P frobG.
+move=> x /setD1P[ntx Kx]; have [_ /mulG_sub[sKG _] _ tiKH] := sdprodP defG.
+have [/eqP <- _ _] := and3P Frobenius_partition; rewrite big_distrl /=.
+apply/bigcupsP=> _ /setU1P[|/imsetP[y Ky]] ->; first exact: subsetIl.
+apply: contraR ntx => /subsetPn[z]; rewrite inE mem_conjg => /andP[Hzy cxz] _.
+rewrite -(conjg_eq1 x y^-1) -in_set1 -set1gE -tiKH inE andbC.
+rewrite -(tiHG _ _ Hzy) ?(subsetP sKG) ?in_group // Ky andbT -conjJg.
+by rewrite /(z ^ x) (cent1P cxz) mulKg.
+Qed.
+
+Lemma Frobenius_reg_ker : semiregular K H.
+Proof.
+move=> x /setD1P[ntx Hx].
+apply/trivgP/subsetP=> y /setIP[Ky cxy]; apply: contraR ntx => nty.
+have K1y: y \in K^# by rewrite inE nty.
+have [/eqP/sdprod_context[_ sHG _ _ tiKH] _] := andP frobG.
+suffices: x \in K :&: H by rewrite tiKH inE.
+by rewrite inE (subsetP (Frobenius_cent1_ker K1y)) // inE cent1C (subsetP sHG).
+Qed.
+
+Lemma Frobenius_reg_compl : semiregular H K.
+Proof. by apply: semiregular_sym; exact: Frobenius_reg_ker. Qed.
+
+Lemma Frobenius_dvd_ker1 : #|H| %| #|K|.-1.
+Proof.
+apply: regular_norm_dvd_pred Frobenius_reg_ker.
+by have[/sdprodP[]] := Frobenius_context.
+Qed.
+
+Lemma ltn_odd_Frobenius_ker : odd #|G| -> #|H|.*2 < #|K|.
+Proof.
+move/oddSg=> oddG.
+have [/sdprodW/mulG_sub[sKG sHG] ntK _ _ _] := Frobenius_context.
+by rewrite dvdn_double_ltn ?oddG ?cardG_gt1 ?Frobenius_dvd_ker1.
+Qed.
+
+Lemma Frobenius_index_dvd_ker1 : #|G : K| %| #|K|.-1.
+Proof.
+have[defG _ _ /andP[sKG _] _] := Frobenius_context.
+by rewrite -divgS // -(sdprod_card defG) mulKn ?Frobenius_dvd_ker1.
+Qed.
+
+Lemma Frobenius_coprime : coprime #|K| #|H|.
+Proof. by rewrite (coprime_dvdr Frobenius_dvd_ker1) ?coprimenP. Qed.
+
+Lemma Frobenius_trivg_cent : 'C_K(H) = 1.
+Proof.
+by apply: (cent_semiregular Frobenius_reg_ker); case: Frobenius_context.
+Qed.
+
+Lemma Frobenius_index_coprime : coprime #|K| #|G : K|.
+Proof. by rewrite (coprime_dvdr Frobenius_index_dvd_ker1) ?coprimenP. Qed.
+
+Lemma Frobenius_ker_Hall : Hall G K.
+Proof.
+have [_ _ _ /andP[sKG _] _] := Frobenius_context.
+by rewrite /Hall sKG Frobenius_index_coprime.
+Qed.
+
+Lemma Frobenius_compl_Hall : Hall G H.
+Proof.
+have [defG _ _ _ _] := Frobenius_context.
+by rewrite -(sdprod_Hall defG) Frobenius_ker_Hall.
+Qed.
+
+End FrobeniusProperties.
+
+Lemma normedTI_J x A G L : normedTI (A :^ x) (G :^ x) (L :^ x) = normedTI A G L.
+Proof.
+rewrite {1}/normedTI normJ -conjIg -(conj0g x) !(can_eq (conjsgK x)).
+congr [&& _, _ == _ & _]; rewrite /cover (reindex_inj (@conjsg_inj _ x)).
+ by apply: eq_big => Hy; rewrite ?orbit_conjsg ?cardJg.
+by rewrite bigcupJ cardJg (eq_bigl _ _ (orbit_conjsg _ _ _ _)).
+Qed.
+
+Lemma FrobeniusJcompl x G H :
+ [Frobenius G :^ x with complement H :^ x] = [Frobenius G with complement H].
+Proof.
+by congr (_ && _); rewrite ?(can_eq (conjsgK x)) // -conjD1g normedTI_J.
+Qed.
+
+Lemma FrobeniusJ x G K H :
+ [Frobenius G :^ x = K :^ x ><| H :^ x] = [Frobenius G = K ><| H].
+Proof.
+by congr (_ && _); rewrite ?FrobeniusJcompl // -sdprodJ (can_eq (conjsgK x)).
+Qed.
+
+Lemma FrobeniusJker x G K :
+ [Frobenius G :^ x with kernel K :^ x] = [Frobenius G with kernel K].
+Proof.
+apply/existsP/existsP=> [] [H]; last by exists (H :^ x)%G; rewrite FrobeniusJ.
+by rewrite -(conjsgKV x H) FrobeniusJ; exists (H :^ x^-1)%G.
+Qed.
+
+Lemma FrobeniusJgroup x G : [Frobenius G :^ x] = [Frobenius G].
+Proof.
+apply/existsP/existsP=> [] [H].
+ by rewrite -(conjsgKV x H) FrobeniusJcompl; exists (H :^ x^-1)%G.
+by exists (H :^ x)%G; rewrite FrobeniusJcompl.
+Qed.
+
+Lemma Frobenius_ker_dvd_ker1 G K :
+ [Frobenius G with kernel K] -> #|G : K| %| #|K|.-1.
+Proof. case/existsP=> H; exact: Frobenius_index_dvd_ker1. Qed.
+
+Lemma Frobenius_ker_coprime G K :
+ [Frobenius G with kernel K] -> coprime #|K| #|G : K|.
+Proof. case/existsP=> H; exact: Frobenius_index_coprime. Qed.
+
+Lemma Frobenius_semiregularP G K H :
+ K ><| H = G -> K :!=: 1 -> H :!=: 1 ->
+ reflect (semiregular K H) [Frobenius G = K ><| H].
+Proof.
+move=> defG ntK ntH.
+apply: (iffP idP) => [|regG]; first exact: Frobenius_reg_ker.
+have [nsKG sHG defKH nKH tiKH]:= sdprod_context defG; have [sKG _]:= andP nsKG.
+apply/and3P; split; first by rewrite defG.
+ by rewrite eqEcard sHG -(sdprod_card defG) -ltnNge ltn_Pmull ?cardG_gt1.
+apply/normedTI_memJ_P; rewrite setD_eq0 subG1 sHG -defKH -(normC nKH).
+split=> // z _ /setD1P[ntz Hz] /mulsgP[y x Hy Kx ->]; rewrite groupMl // !inE.
+rewrite conjg_eq1 ntz; apply/idP/idP=> [Hzxy | Hx]; last by rewrite !in_group.
+apply: (subsetP (sub1G H)); have Hzy: z ^ y \in H by apply: groupJ.
+rewrite -(regG (z ^ y)); last by apply/setD1P; rewrite conjg_eq1.
+rewrite inE Kx cent1C (sameP cent1P commgP) -in_set1 -[[set 1]]tiKH inE /=.
+rewrite andbC groupM ?groupV -?conjgM //= commgEr groupMr //.
+by rewrite memJ_norm ?(subsetP nKH) ?groupV.
+Qed.
+
+Lemma prime_FrobeniusP G K H :
+ K :!=: 1 -> prime #|H| ->
+ reflect (K ><| H = G /\ 'C_K(H) = 1) [Frobenius G = K ><| H].
+Proof.
+move=> ntK H_pr; have ntH: H :!=: 1 by rewrite -cardG_gt1 prime_gt1.
+have [defG | not_sdG] := eqVneq (K ><| H) G; last first.
+ by apply: (iffP andP) => [] [defG]; rewrite defG ?eqxx in not_sdG.
+apply: (iffP (Frobenius_semiregularP defG ntK ntH)) => [regH | [_ regH x]].
+ split=> //; have [x defH] := cyclicP (prime_cyclic H_pr).
+ by rewrite defH cent_cycle regH // !inE defH cycle_id andbT -cycle_eq1 -defH.
+case/setD1P=> nt_x Hx; apply/trivgP; rewrite -regH setIS //= -cent_cycle.
+by rewrite centS // prime_meetG // (setIidPr _) ?cycle_eq1 ?cycle_subG.
+Qed.
+
+Lemma Frobenius_subl G K K1 H :
+ K1 :!=: 1 -> K1 \subset K -> H \subset 'N(K1) -> [Frobenius G = K ><| H] ->
+ [Frobenius K1 <*> H = K1 ><| H].
+Proof.
+move=> ntK1 sK1K nK1H frobG; have [_ _ ntH _ _] := Frobenius_context frobG.
+apply/Frobenius_semiregularP=> //.
+ by rewrite sdprodEY ?coprime_TIg ?(coprimeSg sK1K) ?(Frobenius_coprime frobG).
+by move=> x /(Frobenius_reg_ker frobG) cKx1; apply/trivgP; rewrite -cKx1 setSI.
+Qed.
+
+Lemma Frobenius_subr G K H H1 :
+ H1 :!=: 1 -> H1 \subset H -> [Frobenius G = K ><| H] ->
+ [Frobenius K <*> H1 = K ><| H1].
+Proof.
+move=> ntH1 sH1H frobG; have [defG ntK _ _ _] := Frobenius_context frobG.
+apply/Frobenius_semiregularP=> //.
+ have [_ _ /(subset_trans sH1H) nH1K tiHK] := sdprodP defG.
+ by rewrite sdprodEY //; apply/trivgP; rewrite -tiHK setIS.
+by apply: sub_in1 (Frobenius_reg_ker frobG); exact/subsetP/setSD.
+Qed.
+
+Lemma Frobenius_kerP G K :
+ reflect [/\ K :!=: 1, K \proper G, K <| G
+ & {in K^#, forall x, 'C_G[x] \subset K}]
+ [Frobenius G with kernel K].
+Proof.
+apply: (iffP existsP) => [[H frobG] | [ntK ltKG nsKG regK]].
+ have [/sdprod_context[nsKG _ _ _ _] ntK _ ltKG _] := Frobenius_context frobG.
+ by split=> //; exact: Frobenius_cent1_ker frobG.
+have /andP[sKG nKG] := nsKG.
+have hallK: Hall G K.
+ rewrite /Hall sKG //= coprime_sym coprime_pi' //.
+ apply: sub_pgroup (pgroup_pi K) => p; have [P sylP] := Sylow_exists p G.
+ have [[sPG pP p'GiP] sylPK] := (and3P sylP, Hall_setI_normal nsKG sylP).
+ rewrite -p_rank_gt0 -(rank_Sylow sylPK) rank_gt0 => ntPK.
+ rewrite inE /= -p'natEpi // (pnat_dvd _ p'GiP) ?indexgS //.
+ have /trivgPn[z]: P :&: K :&: 'Z(P) != 1.
+ by rewrite meet_center_nil ?(pgroup_nil pP) ?(normalGI sPG nsKG).
+ rewrite !inE -andbA -sub_cent1=> /and4P[_ Kz _ cPz] ntz.
+ by apply: subset_trans (regK z _); [exact/subsetIP | exact/setD1P].
+have /splitsP[H /complP[tiKH defG]] := SchurZassenhaus_split hallK nsKG.
+have [_ sHG] := mulG_sub defG; have nKH := subset_trans sHG nKG.
+exists H; apply/Frobenius_semiregularP; rewrite ?sdprodE //.
+ by apply: contraNneq (proper_subn ltKG) => H1; rewrite -defG H1 mulg1.
+apply: semiregular_sym => x Kx; apply/trivgP; rewrite -tiKH.
+by rewrite subsetI subsetIl (subset_trans _ (regK x _)) ?setSI.
+Qed.
+
+Lemma set_Frobenius_compl G K H :
+ K ><| H = G -> [Frobenius G with kernel K] -> [Frobenius G = K ><| H].
+Proof.
+move=> defG /Frobenius_kerP[ntK ltKG _ regKG].
+apply/Frobenius_semiregularP=> //.
+ by apply: contraTneq ltKG => H_1; rewrite -defG H_1 sdprodg1 properxx.
+apply: semiregular_sym => y /regKG sCyK.
+have [_ sHG _ _ tiKH] := sdprod_context defG.
+by apply/trivgP; rewrite /= -(setIidPr sHG) setIAC -tiKH setSI.
+Qed.
+
+Lemma Frobenius_kerS G K G1 :
+ G1 \subset G -> K \proper G1 ->
+ [Frobenius G with kernel K] -> [Frobenius G1 with kernel K].
+Proof.
+move=> sG1G ltKG1 /Frobenius_kerP[ntK _ /andP[_ nKG] regKG].
+apply/Frobenius_kerP; rewrite /normal proper_sub // (subset_trans sG1G) //.
+by split=> // x /regKG; apply: subset_trans; rewrite setSI.
+Qed.
+
+Lemma Frobenius_action_kernel_def G H K sT S to :
+ K ><| H = G -> @Frobenius_action _ G H sT S to ->
+ K :=: 1 :|: [set x in G | 'Fix_(S | to)[x] == set0].
+Proof.
+move=> defG FrobG.
+have partG: partition (gval K |: (H^# :^: K)) G.
+ apply: Frobenius_partition; apply/andP; rewrite defG; split=> //.
+ by apply/Frobenius_actionP; exact: HasFrobeniusAction FrobG.
+have{FrobG} [ffulG transG regG ntH [u Su defH]]:= FrobG.
+apply/setP=> x; rewrite !inE; have [-> | ntx] := altP eqP; first exact: group1.
+rewrite /= -(cover_partition partG) /cover.
+have neKHy y: gval K <> H^# :^ y.
+ by move/setP/(_ 1); rewrite group1 conjD1g setD11.
+rewrite big_setU1 /= ?inE; last by apply/imsetP=> [[y _ /neKHy]].
+have [nsKG sHG _ _ tiKH] := sdprod_context defG; have [sKG nKG]:= andP nsKG.
+symmetry; case Kx: (x \in K) => /=.
+ apply/set0Pn=> [[v /setIP[Sv]]]; have [y Gy ->] := atransP2 transG Su Sv.
+ rewrite -sub1set -astabC sub1set astab1_act mem_conjg => Hxy.
+ case/negP: ntx; rewrite -in_set1 -(conjgKV y x) -mem_conjgV conjs1g -tiKH.
+ by rewrite defH setIA inE -mem_conjg (setIidPl sKG) (normsP nKG) ?Kx.
+apply/andP=> [[/bigcupP[_ /imsetP[y Ky ->] Hyx] /set0Pn[]]]; exists (to u y).
+rewrite inE (actsP (atrans_acts transG)) ?(subsetP sKG) // Su.
+rewrite -sub1set -astabC sub1set astab1_act.
+by rewrite conjD1g defH conjIg !inE in Hyx; case/and3P: Hyx.
+Qed.
+
+End FrobeniusBasics.
+
+Implicit Arguments normedTI_P [gT A G L].
+Implicit Arguments normedTI_memJ_P [gT A G L].
+Implicit Arguments Frobenius_kerP [gT G K].
+
+Lemma Frobenius_coprime_quotient (gT : finGroupType) (G K H N : {group gT}) :
+ K ><| H = G -> N <| G -> coprime #|K| #|H| /\ H :!=: 1%g ->
+ N \proper K /\ {in H^#, forall x, 'C_K[x] \subset N} ->
+ [Frobenius G / N = (K / N) ><| (H / N)]%g.
+Proof.
+move=> defG nsNG [coKH ntH] [ltNK regH].
+have [[sNK _] [_ /mulG_sub[sKG sHG] _ _]] := (andP ltNK, sdprodP defG).
+have [_ nNG] := andP nsNG; have nNH := subset_trans sHG nNG.
+apply/Frobenius_semiregularP; first exact: quotient_coprime_sdprod.
+- by rewrite quotient_neq1 ?(normalS _ sKG).
+- by rewrite -(isog_eq1 (quotient_isog _ _)) ?coprime_TIg ?(coprimeSg sNK).
+move=> _ /(subsetP (quotientD1 _ _))/morphimP[x nNx H1x ->].
+rewrite -cent_cycle -quotient_cycle //=.
+rewrite -strongest_coprime_quotient_cent ?cycle_subG //.
+- by rewrite cent_cycle quotientS1 ?regH.
+- by rewrite subIset ?sNK.
+- rewrite (coprimeSg (subsetIl N _)) ?(coprimeSg sNK) ?(coprimegS _ coKH) //.
+ by rewrite cycle_subG; case/setD1P: H1x.
+by rewrite orbC abelian_sol ?cycle_abelian.
+Qed.
+
+Section InjmFrobenius.
+
+Variables (gT rT : finGroupType) (D G : {group gT}) (f : {morphism D >-> rT}).
+Implicit Types (H K : {group gT}) (sGD : G \subset D) (injf : 'injm f).
+
+Lemma injm_Frobenius_compl H sGD injf :
+ [Frobenius G with complement H] -> [Frobenius f @* G with complement f @* H].
+Proof.
+case/andP=> neqGH /normedTI_P[nzH /subsetIP[sHG _] tiHG].
+have sHD := subset_trans sHG sGD; have sH1D := subset_trans (subD1set H 1) sHD.
+apply/andP; rewrite (can_in_eq (injmK injf)) //; split=> //.
+apply/normedTI_P; rewrite normD1 -injmD1 // -!cards_eq0 card_injm // in nzH *.
+rewrite subsetI normG morphimS //; split=> // _ /morphimP[x Dx Gx ->] ti'fHx.
+rewrite mem_morphim ?tiHG //; apply: contra ti'fHx; rewrite -!setI_eq0 => tiHx.
+by rewrite -morphimJ // -injmI ?conj_subG // (eqP tiHx) morphim0.
+Qed.
+
+Lemma injm_Frobenius H K sGD injf :
+ [Frobenius G = K ><| H] -> [Frobenius f @* G = f @* K ><| f @* H].
+Proof.
+case/andP=> /eqP defG frobG.
+by apply/andP; rewrite (injm_sdprod _ injf defG) // eqxx injm_Frobenius_compl.
+Qed.
+
+Lemma injm_Frobenius_ker K sGD injf :
+ [Frobenius G with kernel K] -> [Frobenius f @* G with kernel f @* K].
+Proof.
+case/existsP=> H frobG; apply/existsP; exists (f @* H)%G; exact: injm_Frobenius.
+Qed.
+
+Lemma injm_Frobenius_group sGD injf : [Frobenius G] -> [Frobenius f @* G].
+Proof.
+case/existsP=> H frobG; apply/existsP; exists (f @* H)%G.
+exact: injm_Frobenius_compl.
+Qed.
+
+End InjmFrobenius.
+
+Theorem Frobenius_Ldiv (gT : finGroupType) (G : {group gT}) n :
+ n %| #|G| -> n %| #|'Ldiv_n(G)|.
+Proof.
+move=> nG; move: {2}_.+1 (ltnSn (#|G| %/ n)) => mq.
+elim: mq => // mq IHm in gT G n nG *; case/dvdnP: nG => q oG.
+have [q_gt0 n_gt0] : 0 < q /\ 0 < n by apply/andP; rewrite -muln_gt0 -oG.
+rewrite ltnS oG mulnK // => leqm.
+have:= q_gt0; rewrite leq_eqVlt => /predU1P[q1 | lt1q].
+ rewrite -(mul1n n) q1 -oG (setIidPl _) //.
+ by apply/subsetP=> x Gx; rewrite inE -order_dvdn order_dvdG.
+pose p := pdiv q; have pr_p: prime p by exact: pdiv_prime.
+have lt1p: 1 < p := prime_gt1 pr_p; have p_gt0 := ltnW lt1p.
+have{leqm} lt_qp_mq: q %/ p < mq by apply: leq_trans leqm; rewrite ltn_Pdiv.
+have: n %| #|'Ldiv_(p * n)(G)|.
+ have: p * n %| #|G| by rewrite oG dvdn_pmul2r ?pdiv_dvd.
+ move/IHm=> IH; apply: dvdn_trans (IH _); first exact: dvdn_mull.
+ by rewrite oG divnMr.
+rewrite -(cardsID 'Ldiv_n()) dvdn_addl.
+ rewrite -setIA ['Ldiv_n(_)](setIidPr _) //.
+ by apply/subsetP=> x; rewrite !inE -!order_dvdn; apply: dvdn_mull.
+rewrite -setIDA; set A := _ :\: _.
+have pA x: x \in A -> #[x]`_p = (n`_p * p)%N.
+ rewrite !inE -!order_dvdn => /andP[xn xnp].
+ rewrite !p_part // -expnSr; congr (p ^ _)%N; apply/eqP.
+ rewrite eqn_leq -{1}addn1 -(pfactorK 1 pr_p) -lognM ?expn1 // mulnC.
+ rewrite dvdn_leq_log ?muln_gt0 ?p_gt0 //= ltnNge; apply: contra xn => xn.
+ move: xnp; rewrite -[#[x]](partnC p) //.
+ rewrite !Gauss_dvd ?coprime_partC //; case/andP=> _.
+ rewrite p_part ?pfactor_dvdn // xn Gauss_dvdr // coprime_sym.
+ exact: pnat_coprime (pnat_id _) (part_pnat _ _).
+rewrite -(partnC p n_gt0) Gauss_dvd ?coprime_partC //; apply/andP; split.
+ rewrite -sum1_card (partition_big_imset (@cycle _)) /=.
+ apply: dvdn_sum => _ /imsetP[x /setIP[Gx Ax] ->].
+ rewrite (eq_bigl (generator <[x]>)) => [|y].
+ rewrite sum1dep_card -totient_gen -[#[x]](partnC p) //.
+ rewrite totient_coprime ?coprime_partC // dvdn_mulr // .
+ by rewrite (pA x Ax) p_part // -expnSr totient_pfactor // dvdn_mull.
+ rewrite /generator eq_sym andbC; case xy: {+}(_ == _) => //.
+ rewrite !inE -!order_dvdn in Ax *.
+ by rewrite -cycle_subG /order -(eqP xy) cycle_subG Gx.
+rewrite -sum1_card (partition_big_imset (fun x => x.`_p ^: G)) /=.
+apply: dvdn_sum => _ /imsetP[x /setIP[Gx Ax] ->].
+set y := x.`_p; have oy: #[y] = (n`_p * p)%N by rewrite order_constt pA.
+rewrite (partition_big (fun x => x.`_p) (mem (y ^: G))) /= => [|z]; last first.
+ by case/andP=> _ /eqP <-; rewrite /= class_refl.
+pose G' := ('C_G[y] / <[y]>)%G; pose n' := gcdn #|G'| n`_p^'.
+have n'_gt0: 0 < n' by rewrite gcdn_gt0 cardG_gt0.
+rewrite (eq_bigr (fun _ => #|'Ldiv_n'(G')|)) => [|_ /imsetP[a Ga ->]].
+ rewrite sum_nat_const -index_cent1 indexgI.
+ rewrite -(dvdn_pmul2l (cardG_gt0 'C_G[y])) mulnA LagrangeI.
+ have oCy: #|'C_G[y]| = (#[y] * #|G'|)%N.
+ rewrite card_quotient ?subcent1_cycle_norm // Lagrange //.
+ by rewrite subcent1_cycle_sub ?groupX.
+ rewrite oCy -mulnA -(muln_lcm_gcd #|G'|) -/n' mulnA dvdn_mul //.
+ rewrite muln_lcmr -oCy order_constt pA // mulnAC partnC // dvdn_lcm.
+ by rewrite cardSg ?subsetIl // mulnC oG dvdn_pmul2r ?pdiv_dvd.
+ apply: IHm; [exact: dvdn_gcdl | apply: leq_ltn_trans lt_qp_mq].
+ rewrite -(@divnMr n`_p^') // -muln_lcm_gcd mulnC divnMl //.
+ rewrite leq_divRL // divn_mulAC ?leq_divLR ?dvdn_mulr ?dvdn_lcmr //.
+ rewrite dvdn_leq ?muln_gt0 ?q_gt0 //= mulnC muln_lcmr dvdn_lcm.
+ rewrite -(@dvdn_pmul2l n`_p) // mulnA -oy -oCy mulnCA partnC // -oG.
+ by rewrite cardSg ?subsetIl // dvdn_mul ?pdiv_dvd.
+pose h := [fun z => coset <[y]> (z ^ a^-1)].
+pose h' := [fun Z : coset_of <[y]> => (y * (repr Z).`_p^') ^ a].
+rewrite -sum1_card (reindex_onto h h') /= => [|Z]; last first.
+ rewrite conjgK coset_kerl ?cycle_id ?morph_constt ?repr_coset_norm //.
+ rewrite /= coset_reprK 2!inE -order_dvdn dvdn_gcd => /and3P[_ _ p'Z].
+ by apply: constt_p_elt (pnat_dvd p'Z _); apply: part_pnat.
+apply: eq_bigl => z; apply/andP/andP=> [[]|[]].
+ rewrite inE -andbA => /and3P[Gz Az _] /eqP zp_ya.
+ have czy: z ^ a^-1 \in 'C[y].
+ rewrite -mem_conjg -normJ conjg_set1 -zp_ya.
+ by apply/cent1P; apply: commuteX.
+ have Nz: z ^ a^-1 \in 'N(<[y]>) by apply: subsetP czy; apply: norm_gen.
+ have G'z: h z \in G' by rewrite mem_morphim //= inE groupJ // groupV.
+ rewrite inE G'z inE -order_dvdn dvdn_gcd order_dvdG //=.
+ rewrite /order -morphim_cycle // -quotientE card_quotient ?cycle_subG //.
+ rewrite -(@dvdn_pmul2l #[y]) // Lagrange; last first.
+ by rewrite /= cycleJ cycle_subG mem_conjgV -zp_ya mem_cycle.
+ rewrite oy mulnAC partnC // [#|_|]orderJ; split.
+ by rewrite !inE -!order_dvdn mulnC in Az; case/andP: Az.
+ set Z := coset _ _; have NZ := repr_coset_norm Z; have:= coset_reprK Z.
+ case/kercoset_rcoset=> {NZ}// _ /cycleP[i ->] ->{Z}.
+ rewrite consttM; last exact/commute_sym/commuteX/cent1P.
+ rewrite (constt1P _) ?p_eltNK 1?p_eltX ?p_elt_constt // mul1g.
+ by rewrite conjMg consttJ conjgKV -zp_ya consttC.
+rewrite 2!inE -order_dvdn; set Z := coset _ _ => /andP[Cz n'Z] /eqP def_z.
+have Nz: z ^ a^-1 \in 'N(<[y]>).
+ rewrite -def_z conjgK groupMr; first by rewrite -(cycle_subG y) normG.
+ by rewrite groupX ?repr_coset_norm.
+have{Cz} /setIP[Gz Cz]: z ^ a^-1 \in 'C_G[y].
+ case/morphimP: Cz => u Nu Cu /kercoset_rcoset[] // _ /cycleP[i ->] ->.
+ by rewrite groupMr // groupX // inE groupX //; apply/cent1P.
+have{def_z} zp_ya: z.`_p = y ^ a.
+ rewrite -def_z consttJ consttM.
+ rewrite constt_p_elt ?p_elt_constt //.
+ by rewrite (constt1P _) ?p_eltNK ?p_elt_constt ?mulg1.
+ apply: commute_sym; apply/cent1P.
+ by rewrite -def_z conjgK groupMl // in Cz; apply/cent1P.
+have ozp: #[z ^ a^-1]`_p = #[y] by rewrite -order_constt consttJ zp_ya conjgK.
+split; rewrite zp_ya // -class_lcoset lcoset_id // eqxx andbT.
+rewrite -(conjgKV a z) !inE groupJ //= -!order_dvdn orderJ; apply/andP; split.
+ apply: contra (partn_dvd p n_gt0) _.
+ by rewrite ozp -(muln1 n`_p) oy dvdn_pmul2l // dvdn1 neq_ltn lt1p orbT.
+rewrite -(partnC p n_gt0) mulnCA mulnA -oy -(@partnC p #[_]) // ozp.
+apply dvdn_mul => //; apply: dvdn_trans (dvdn_trans n'Z (dvdn_gcdr _ _)).
+rewrite {2}/order -morphim_cycle // -quotientE card_quotient ?cycle_subG //.
+rewrite -(@dvdn_pmul2l #|<[z ^ a^-1]> :&: <[y]>|) ?cardG_gt0 // LagrangeI.
+rewrite -[#|<[_]>|](partnC p) ?order_gt0 // dvdn_pmul2r // ozp.
+by rewrite cardSg ?subsetIr.
+Qed.
diff --git a/mathcomp/solvable/gfunctor.v b/mathcomp/solvable/gfunctor.v
new file mode 100644
index 0000000..a338999
--- /dev/null
+++ b/mathcomp/solvable/gfunctor.v
@@ -0,0 +1,484 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat fintype bigop finset.
+Require Import fingroup morphism automorphism quotient gproduct.
+
+(******************************************************************************)
+(* This file provides basic interfaces for the notion of "generic" *)
+(* characteristic subgroups; these amount to subfunctors of the identity *)
+(* functor in some category of groups. *)
+(* See "Generic Proof Tools And Finite Group Theory", *)
+(* Francois Garillot, PhD, 2011, Chapter 3. *)
+(* The implementation proposed here is fairly basic, relying on first order *)
+(* function matching and on structure telescopes, both of which are somewhat *)
+(* limited and fragile. It should switch in the future to more general and *)
+(* more robust quotation matching. *)
+(* The definitions in this file (types, properties and structures) are all *)
+(* packaged under the GFunctor submodule, i.e., client code should refer to *)
+(* GFunctor.continuous, GFunctor.map, etc. Notations, Coercions and Lemmas *)
+(* are exported and thus directly available, however. *)
+(* We provide the following: *)
+(* object_map == the type of the (polymorphic) object map of a group *)
+(* functor; the %gF scope is bound to object_map. *)
+(* := forall gT : finGroupType, {set gT} -> {set gT}. *)
+(* We define two operations on object_map (with notations in the %gF scope): *)
+(* F1 \o F2 == the composite map; (F1 \o F2) G expands to F1 (F2 G). *)
+(* F1 %% F2 == F1 computed modulo F2; we have *)
+(* (F1 %% F2) G / F2 G = F1 (G / F2 G) *)
+(* We define the following (type-polymorphic) properties of an object_map F: *)
+(* group_valued F <-> F G is a group when G is a group *)
+(* closed F <-> F G is a subgroup o fG when G is a group *)
+(* continuous F <-> F is continuous with respect to morphism image: *)
+(* for any f : {morphism G >-> ..}, f @* (F G) is a *)
+(* a subgroup of F (f @* G); equivalently, F is *)
+(* functorial in the category Grp of groups. *)
+(* Most common "characteristic subgroup" are produced *)
+(* continuous object maps. *)
+(* iso_continuous F <-> F is continuous with respect to isomorphism image; *)
+(* equivalently, F is functorial in the Grp groupoid. *)
+(* The Puig and the Thompson J subgroups are examples *)
+(* of iso_continuous maps that are not continuous. *)
+(* pcontinuous F <-> F is continuous with respect to partial morphism *)
+(* image, i.e., functorial in the category of groups *)
+(* and partial morphisms. The center and p-core are *)
+(* examples of pcontinuous maps. *)
+(* hereditary F <-> inclusion in the image of F is hereditary, i.e., *)
+(* for any subgroup H of G, the intersection of H with *)
+(* F G is included in H. Note that F is pcontinuous *)
+(* iff it is continuous and hereditary; indeed proofs *)
+(* of pcontinuous F coerce to proofs of hereditary F *)
+(* and continuous F. *)
+(* monotonic F <-> F is monotonic with respect to inclusion: for any *)
+(* subgroup H of G, F H is a subgroup of F G. The *)
+(* derived and lower central series are examples of *)
+(* monotonic maps. *)
+(* Four structures provide interfaces to these properties: *)
+(* GFunctor.iso_map == structure for object maps that are group_valued, *)
+(* closed, and iso_continuous. *)
+(* [igFun by Fsub & !Fcont] == the iso_map structure for an object map F *)
+(* such that F G is canonically a group when G is, and *)
+(* given Fsub : closed F and Fcont : iso_continuous F. *)
+(* [igFun by Fsub & Fcont] == as above, but expecting Fcont : continuous F. *)
+(* [igFun of F] == clone an existing GFunctor.iso_map structure for F. *)
+(* GFunctor.map == structure for continuous object maps, inheriting *)
+(* from the GFunctor.iso_map structure. *)
+(* [gFun by Fcont] == the map structure for an F with a canonical iso_map *)
+(* structure, given Fcont : continuous F. *)
+(* [gFun of F] == clone an existing GFunctor.map structure for F. *)
+(* GFunctor.pmap == structure for pcontinuous object maps, inheriting *)
+(* from the GFunctor.map structure. *)
+(* [pgFun by Fher] == the pmap structure for an F with a canonical map *)
+(* structure, given Fher : hereditary F. *)
+(* [pgFun of F] == clone an existing GFunctor.pmap structure for F. *)
+(* GFunctor.mono_map == structure for monotonic, continuous object maps *)
+(* inheriting from the GFunctor.map structure. *)
+(* [mgFun by Fmon] == the mono_map structure for an F with a canonical *)
+(* map structure, given Fmon : monotonic F. *)
+(* [mgFun of F] == clone an existing GFunctor.mono_map structure for F *)
+(* Lemmas for these group functors use either a 'gF' prefix or an 'F' suffix. *)
+(* The (F1 \o F2) and (F1 %% F2) operations have canonical GFunctor.map *)
+(* structures when F1 is monotonic or hereditary, respectively. *)
+(******************************************************************************)
+
+Import GroupScope.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Delimit Scope gFun_scope with gF.
+
+Module GFunctor.
+
+Definition object_map := forall gT : finGroupType, {set gT} -> {set gT}.
+
+Bind Scope gFun_scope with object_map.
+
+Section Definitions.
+
+Implicit Types gT hT : finGroupType.
+
+Variable F : object_map.
+
+(* Group closure. *)
+Definition group_valued := forall gT (G : {group gT}), group_set (F G).
+
+(* Subgroup closure. *)
+Definition closed := forall gT (G : {group gT}), F G \subset G.
+
+(* General functoriality, i.e., continuity of the object map *)
+Definition continuous :=
+ forall gT hT (G : {group gT}) (phi : {morphism G >-> hT}),
+ phi @* F G \subset F (phi @* G).
+
+(* Functoriality on the Grp groupoid (arrows are restricted to isos). *)
+Definition iso_continuous :=
+ forall gT hT (G : {group gT}) (phi : {morphism G >-> hT}),
+ 'injm phi -> phi @* F G \subset F (phi @* G).
+
+Lemma continuous_is_iso_continuous : continuous -> iso_continuous.
+Proof. by move=> Fcont gT hT G phi inj_phi; exact: Fcont. Qed.
+
+(* Functoriality on Grp with partial morphisms. *)
+Definition pcontinuous :=
+ forall gT hT (G D : {group gT}) (phi : {morphism D >-> hT}),
+ phi @* F G \subset F (phi @* G).
+
+Lemma pcontinuous_is_continuous : pcontinuous -> continuous.
+Proof. by move=> Fcont gT hT G; exact: Fcont. Qed.
+
+(* Heredity with respect to inclusion *)
+Definition hereditary :=
+ forall gT (H G : {group gT}), H \subset G -> F G :&: H \subset F H.
+
+Lemma pcontinuous_is_hereditary : pcontinuous -> hereditary.
+Proof.
+move=> Fcont gT H G sHG; rewrite -{2}(setIidPl sHG) setIC.
+by do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom ?Fcont.
+Qed.
+
+(* Monotonicity with respect to inclusion *)
+Definition monotonic :=
+ forall gT (H G : {group gT}), H \subset G -> F H \subset F G.
+
+(* Self-expanding composition, and modulo *)
+
+Variables (k : unit) (F1 F2 : object_map).
+
+Definition comp_head : object_map := fun gT A => let: tt := k in F1 (F2 A).
+
+Definition modulo : object_map :=
+ fun gT A => coset (F2 A) @*^-1 (F1 (A / (F2 A))).
+
+End Definitions.
+
+Section ClassDefinitions.
+
+Structure iso_map := IsoMap {
+ apply : object_map;
+ _ : group_valued apply;
+ _ : closed apply;
+ _ : iso_continuous apply
+}.
+Local Coercion apply : iso_map >-> object_map.
+
+Structure map := Map { iso_of_map : iso_map; _ : continuous iso_of_map }.
+Local Coercion iso_of_map : map >-> iso_map.
+
+Structure pmap := Pmap { map_of_pmap : map; _ : hereditary map_of_pmap }.
+Local Coercion map_of_pmap : pmap >-> map.
+
+Structure mono_map := MonoMap { map_of_mono : map; _ : monotonic map_of_mono }.
+Local Coercion map_of_mono : mono_map >-> map.
+
+Definition pack_iso F Fcont Fgrp Fsub := @IsoMap F Fgrp Fsub Fcont.
+
+Definition clone_iso (F : object_map) :=
+ fun Fgrp Fsub Fcont (isoF := @IsoMap F Fgrp Fsub Fcont) =>
+ fun isoF0 & phant_id (apply isoF0) F & phant_id isoF isoF0 => isoF.
+
+Definition clone (F : object_map) :=
+ fun isoF & phant_id (apply isoF) F =>
+ fun (funF0 : map) & phant_id (apply funF0) F =>
+ fun Fcont (funF := @Map isoF Fcont) & phant_id funF0 funF => funF.
+
+Definition clone_pmap (F : object_map) :=
+ fun (funF : map) & phant_id (apply funF) F =>
+ fun (pfunF0 : pmap) & phant_id (apply pfunF0) F =>
+ fun Fher (pfunF := @Pmap funF Fher) & phant_id pfunF0 pfunF => pfunF.
+
+Definition clone_mono (F : object_map) :=
+ fun (funF : map) & phant_id (apply funF) F =>
+ fun (mfunF0 : mono_map) & phant_id (apply mfunF0) F =>
+ fun Fmon (mfunF := @MonoMap funF Fmon) & phant_id mfunF0 mfunF => mfunF.
+
+End ClassDefinitions.
+
+Module Exports.
+
+Identity Coercion fun_of_object_map : object_map >-> Funclass.
+Coercion apply : iso_map >-> object_map.
+Coercion iso_of_map : map >-> iso_map.
+Coercion map_of_pmap : pmap >-> map.
+Coercion map_of_mono : mono_map >-> map.
+Coercion continuous_is_iso_continuous : continuous >-> iso_continuous.
+Coercion pcontinuous_is_continuous : pcontinuous >-> continuous.
+Coercion pcontinuous_is_hereditary : pcontinuous >-> hereditary.
+
+Notation "[ 'igFun' 'by' Fsub & Fcont ]" :=
+ (pack_iso (continuous_is_iso_continuous Fcont) (fun gT G => groupP _) Fsub)
+ (at level 0, format "[ 'igFun' 'by' Fsub & Fcont ]") : form_scope.
+
+Notation "[ 'igFun' 'by' Fsub & ! Fcont ]" :=
+ (pack_iso Fcont (fun gT G => groupP _) Fsub)
+ (at level 0, format "[ 'igFun' 'by' Fsub & ! Fcont ]") : form_scope.
+
+Notation "[ 'igFun' 'of' F ]" := (@clone_iso F _ _ _ _ id id)
+ (at level 0, format "[ 'igFun' 'of' F ]") : form_scope.
+
+Notation "[ 'gFun' 'by' Fcont ]" := (Map Fcont)
+ (at level 0, format "[ 'gFun' 'by' Fcont ]") : form_scope.
+
+Notation "[ 'gFun' 'of' F ]" := (@clone F _ id _ id _ id)
+ (at level 0, format "[ 'gFun' 'of' F ]") : form_scope.
+
+Notation "[ 'pgFun' 'by' Fher ]" := (Pmap Fher)
+ (at level 0, format "[ 'pgFun' 'by' Fher ]") : form_scope.
+
+Notation "[ 'pgFun' 'of' F ]" := (@clone_pmap F _ id _ id _ id)
+ (at level 0, format "[ 'pgFun' 'of' F ]") : form_scope.
+
+Notation "[ 'mgFun' 'by' Fmon ]" := (MonoMap Fmon)
+ (at level 0, format "[ 'mgFun' 'by' Fmon ]") : form_scope.
+
+Notation "[ 'mgFun' 'of' F ]" := (@clone_mono F _ id _ id _ id)
+ (at level 0, format "[ 'mgFun' 'of' F ]") : form_scope.
+
+End Exports.
+
+End GFunctor.
+Export GFunctor.Exports.
+
+Bind Scope gFun_scope with GFunctor.object_map.
+
+Notation "F1 \o F2" := (GFunctor.comp_head tt F1 F2) : gFun_scope.
+Notation "F1 %% F2" := (GFunctor.modulo F1 F2) : gFun_scope.
+
+Section FunctorGroup.
+
+Variables (F : GFunctor.iso_map) (gT : finGroupType) (G : {group gT}).
+Lemma gFgroupset : group_set (F gT G). Proof. by case: F. Qed.
+Canonical gFgroup := Group gFgroupset.
+
+End FunctorGroup.
+
+Canonical gFmod_group
+ (F1 : GFunctor.iso_map) (F2 : GFunctor.object_map)
+ (gT : finGroupType) (G : {group gT}) :=
+ [group of (F1 %% F2)%gF gT G].
+
+Section IsoFunctorTheory.
+
+Implicit Types gT rT : finGroupType.
+Variable F : GFunctor.iso_map.
+
+Lemma gFsub gT (G : {group gT}) : F gT G \subset G.
+Proof. by case: F gT G. Qed.
+
+Lemma gF1 gT : F gT 1 = 1. Proof. exact/trivgP/gFsub. Qed.
+
+Lemma gFiso_cont : GFunctor.iso_continuous F.
+Proof. by case F. Qed.
+
+Lemma gFchar gT (G : {group gT}) : F gT G \char G.
+Proof.
+apply/andP; split => //; first by apply: gFsub.
+apply/forall_inP=> f Af; rewrite -{2}(im_autm Af) -(autmE Af).
+by rewrite -morphimEsub ?gFsub ?gFiso_cont ?injm_autm.
+Qed.
+
+Lemma gFnorm gT (G : {group gT}) : G \subset 'N(F gT G).
+Proof. by rewrite char_norm ?gFchar. Qed.
+
+Lemma gFnormal gT (G : {group gT}) : F gT G <| G.
+Proof. by rewrite char_normal ?gFchar. Qed.
+
+Lemma injmF_sub gT rT (G D : {group gT}) (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> f @* (F gT G) \subset F rT (f @* G).
+Proof.
+move=> injf sGD; apply/eqP; rewrite -(setIidPr (gFsub G)).
+by rewrite-{3}(setIid G) -!(morphim_restrm sGD) gFiso_cont // injm_restrm.
+Qed.
+
+Lemma injmF gT rT (G D : {group gT}) (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> f @* (F gT G) = F rT (f @* G).
+Proof.
+move=> injf sGD; apply/eqP; rewrite eqEsubset injmF_sub //=.
+rewrite -{2}(morphim_invm injf sGD) -[f @* F _ _](morphpre_invm injf).
+have Fsubs := subset_trans (gFsub _).
+by rewrite -sub_morphim_pre (injmF_sub, Fsubs) ?morphimS ?injm_invm.
+Qed.
+
+Lemma gFisom gT rT (G D : {group gT}) R (f : {morphism D >-> rT}) :
+ G \subset D -> isom G (gval R) f -> isom (F gT G) (F rT R) f.
+Proof.
+case/(restrmP f)=> g [gf _ _ _]; rewrite -{f}gf.
+by case/isomP=> injg <-; rewrite sub_isom ?gFsub ?injmF.
+Qed.
+
+Lemma gFisog gT rT (G : {group gT}) (R : {group rT}) :
+ G \isog R -> F gT G \isog F rT R.
+Proof. by case/isogP=> f injf <-; rewrite -injmF // sub_isog ?gFsub. Qed.
+
+End IsoFunctorTheory.
+
+Section FunctorTheory.
+
+Implicit Types gT rT : finGroupType.
+Variable F : GFunctor.map.
+
+Lemma gFcont : GFunctor.continuous F.
+Proof. by case F. Qed.
+
+Lemma morphimF gT rT (G D : {group gT}) (f : {morphism D >-> rT}) :
+ G \subset D -> f @* (F gT G) \subset F rT (f @* G).
+Proof.
+move=> sGD; rewrite -(setIidPr (gFsub F G)).
+by rewrite -{3}(setIid G) -!(morphim_restrm sGD) gFcont.
+Qed.
+
+End FunctorTheory.
+
+Section PartialFunctorTheory.
+
+Implicit Types gT rT : finGroupType.
+
+Section BasicTheory.
+
+Variable F : GFunctor.pmap.
+
+Lemma gFhereditary : GFunctor.hereditary F.
+Proof. by case F. Qed.
+
+Lemma gFunctorI gT (G H : {group gT}) :
+ F gT G :&: H = F gT G :&: F gT (G :&: H).
+Proof.
+rewrite -{1}(setIidPr (gFsub F G)) [G :&: _]setIC -setIA.
+rewrite -(setIidPr (gFhereditary (subsetIl G H))).
+by rewrite setIC -setIA (setIidPr (gFsub F (G :&: H))).
+Qed.
+
+Lemma pmorphimF : GFunctor.pcontinuous F.
+Proof.
+move=> gT rT G D f; rewrite -morphimIdom -(setIidPl (gFsub F G)) setICA.
+apply: (subset_trans (morphimS f (gFhereditary (subsetIr D G)))).
+by rewrite (subset_trans (morphimF F _ _ )) ?morphimIdom ?subsetIl.
+Qed.
+
+Lemma gFid gT (G : {group gT}) : F gT (F gT G) = F gT G.
+Proof.
+apply/eqP; rewrite eqEsubset gFsub.
+by move/gFhereditary: (gFsub F G); rewrite setIid /=.
+Qed.
+
+End BasicTheory.
+
+Section Modulo.
+
+Variables (F1 : GFunctor.pmap) (F2 : GFunctor.map).
+
+Lemma gFmod_closed : GFunctor.closed (F1 %% F2).
+Proof. by move=> gT G; rewrite sub_cosetpre_quo ?gFsub ?gFnormal. Qed.
+
+Lemma gFmod_cont : GFunctor.continuous (F1 %% F2).
+Proof.
+move=> gT rT G f; have nF2 := gFnorm F2.
+have sDF: G \subset 'dom (coset (F2 _ G)) by rewrite nF2.
+have sDFf: G \subset 'dom (coset (F2 _ (f @* G)) \o f).
+ by rewrite -sub_morphim_pre ?subsetIl // nF2.
+pose K := 'ker (restrm sDFf (coset (F2 _ (f @* G)) \o f)).
+have sFK: 'ker (restrm sDF (coset (F2 _ G))) \subset K.
+ rewrite {}/K !ker_restrm ker_comp /= subsetI subsetIl !ker_coset /=.
+ by rewrite -sub_morphim_pre ?subsetIl // morphimIdom ?morphimF.
+have sOF := gFsub F1 (G / F2 _ G); have sGG: G \subset G by [].
+rewrite -sub_quotient_pre; last first.
+ by apply: subset_trans (nF2 _ _); rewrite morphimS ?gFmod_closed.
+suffices im_fact H : F2 _ G \subset gval H -> H \subset G ->
+ factm sFK sGG @* (H / F2 _ G) = f @* H / F2 _ (f @* G).
+- rewrite -2?im_fact ?gFmod_closed ?gFsub //.
+ by rewrite cosetpreK morphimF /= ?morphim_restrm ?setIid.
+ by rewrite -sub_quotient_pre ?normG //= trivg_quotient sub1G.
+move=> sFH sHG; rewrite -(morphimIdom _ (H / _)) /= {2}morphim_restrm setIid.
+rewrite -morphimIG ?ker_coset // -(morphim_restrm sDF) morphim_factm.
+by rewrite morphim_restrm morphim_comp -quotientE morphimIdom.
+Qed.
+
+Canonical gFmod_igFun := [igFun by gFmod_closed & gFmod_cont].
+Canonical gFmod_gFun := [gFun by gFmod_cont].
+
+End Modulo.
+
+Variables F1 F2 : GFunctor.pmap.
+
+Lemma gFmod_hereditary : GFunctor.hereditary (F1 %% F2).
+Proof.
+move=> gT H G sHG; set FGH := _ :&: H; have nF2H := gFnorm F2 H.
+rewrite -sub_quotient_pre; last exact: subset_trans (subsetIr _ _) _.
+pose rH := restrm nF2H (coset (F2 _ H)); pose rHM := [morphism of rH].
+have rnorm_simpl: rHM @* H = H / F2 _ H by rewrite morphim_restrm setIid.
+have nF2G := subset_trans sHG (gFnorm F2 G).
+pose rG := restrm nF2G (coset (F2 _ G)); pose rGM := [morphism of rG].
+have sqKfK: 'ker rGM \subset 'ker rHM.
+ rewrite !ker_restrm !ker_coset (setIidPr (gFsub F2 _)) setIC /=.
+ exact: gFhereditary.
+have sHH := subxx H; rewrite -rnorm_simpl /= -(morphim_factm sqKfK sHH) /=.
+apply: subset_trans (gFcont F1 _); rewrite /= {2}morphim_restrm setIid /=.
+apply: subset_trans (morphimS _ (gFhereditary _ (quotientS _ sHG))) => /=.
+have ->: FGH / _ = restrm nF2H (coset _) @* FGH.
+ by rewrite morphim_restrm setICA setIid.
+rewrite -(morphim_factm sqKfK sHH) morphimS //= morphim_restrm -quotientE.
+by rewrite setICA setIid (subset_trans (quotientI _ _ _)) // cosetpreK.
+Qed.
+
+Canonical gFmod_pgFun := [pgFun by gFmod_hereditary].
+
+End PartialFunctorTheory.
+
+Section MonotonicFunctorTheory.
+
+Implicit Types gT rT : finGroupType.
+
+Lemma gFunctorS (F : GFunctor.mono_map) : GFunctor.monotonic F.
+Proof. by case: F. Qed.
+
+Section Composition.
+
+Variables (F1 : GFunctor.mono_map) (F2 : GFunctor.map).
+
+Lemma gFcomp_closed : GFunctor.closed (F1 \o F2).
+Proof. by move=> gT G; rewrite (subset_trans (gFsub _ _)) ?gFsub. Qed.
+
+Lemma gFcomp_cont : GFunctor.continuous (F1 \o F2).
+Proof.
+move=> gT rT G phi; rewrite (subset_trans (morphimF _ _ (gFsub _ _))) //.
+by rewrite (subset_trans (gFunctorS F1 (gFcont F2 phi))).
+Qed.
+
+Canonical gFcomp_igFun := [igFun by gFcomp_closed & gFcomp_cont].
+Canonical gFcomp_gFun :=[gFun by gFcomp_cont].
+
+End Composition.
+
+Variables F1 F2 : GFunctor.mono_map.
+
+Lemma gFcompS : GFunctor.monotonic (F1 \o F2).
+Proof. by move=> gT H G sHG; rewrite !gFunctorS. Qed.
+
+Canonical gFcomp_mgFun := [mgFun by gFcompS].
+
+End MonotonicFunctorTheory.
+
+Section GFunctorExamples.
+
+Implicit Types gT : finGroupType.
+
+Definition idGfun gT := @id {set gT}.
+
+Lemma idGfun_closed : GFunctor.closed idGfun. Proof. by []. Qed.
+Lemma idGfun_cont : GFunctor.continuous idGfun. Proof. by []. Qed.
+Lemma idGfun_monotonic : GFunctor.monotonic idGfun. Proof. by []. Qed.
+
+Canonical bgFunc_id := [igFun by idGfun_closed & idGfun_cont].
+Canonical gFunc_id := [gFun by idGfun_cont].
+Canonical mgFunc_id := [mgFun by idGfun_monotonic].
+
+Definition trivGfun gT of {set gT} := [1 gT].
+
+Lemma trivGfun_cont : GFunctor.pcontinuous trivGfun.
+Proof. by move=> gT rT D G f; rewrite morphim1. Qed.
+
+Canonical trivGfun_igFun := [igFun by sub1G & trivGfun_cont].
+Canonical trivGfun_gFun := [gFun by trivGfun_cont].
+Canonical trivGfun_pgFun := [pgFun by trivGfun_cont].
+
+End GFunctorExamples.
+
diff --git a/mathcomp/solvable/gseries.v b/mathcomp/solvable/gseries.v
new file mode 100644
index 0000000..4e299e3
--- /dev/null
+++ b/mathcomp/solvable/gseries.v
@@ -0,0 +1,546 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path fintype bigop.
+Require Import finset fingroup morphism automorphism quotient action.
+Require Import commutator center.
+(******************************************************************************)
+(* H <|<| G <=> H is subnormal in G, i.e., H <| ... <| G. *)
+(* invariant_factor A H G <=> A normalises both H and G, and H <| G. *)
+(* A.-invariant <=> the (invariant_factor A) relation, in the context *)
+(* of the g_rel.-series notation. *)
+(* g_rel.-series H s <=> H :: s is a sequence of groups whose projection *)
+(* to sets satisfies relation g_rel pairwise; for *)
+(* example H <|<| G iff G = last H s for some s such *)
+(* that normal.-series H s. *)
+(* stable_factor A H G == H <| G and A centralises G / H. *)
+(* A.-stable == the stable_factor relation, in the scope of the *)
+(* r.-series notation. *)
+(* G.-central == the central_factor relation, in the scope of the *)
+(* r.-series notation. *)
+(* maximal M G == M is a maximal proper subgroup of G. *)
+(* maximal_eq M G == (M == G) or (maximal M G). *)
+(* maxnormal M G N == M is a maximal subgroup of G normalized by N. *)
+(* minnormal M N == M is a minimal nontrivial group normalized by N. *)
+(* simple G == G is a (nontrivial) simple group. *)
+(* := minnormal G G *)
+(* G.-chief == the chief_factor relation, in the scope of the *)
+(* r.-series notation. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section GroupDefs.
+
+Variable gT : finGroupType.
+Implicit Types A B U V : {set gT}.
+
+Notation Local groupT := (group_of (Phant gT)).
+
+Definition subnormal A B :=
+ (A \subset B) && (iter #|B| (fun N => generated (class_support A N)) B == A).
+
+Definition invariant_factor A B C :=
+ [&& A \subset 'N(B), A \subset 'N(C) & B <| C].
+
+Definition group_rel_of (r : rel {set gT}) := [rel H G : groupT | r H G].
+
+Definition stable_factor A V U :=
+ ([~: U, A] \subset V) && (V <| U). (* this orders allows and3P to be used *)
+
+Definition central_factor A V U :=
+ [&& [~: U, A] \subset V, V \subset U & U \subset A].
+
+Definition maximal A B := [max A of G | G \proper B].
+
+Definition maximal_eq A B := (A == B) || maximal A B.
+
+Definition maxnormal A B U := [max A of G | G \proper B & U \subset 'N(G)].
+
+Definition minnormal A B := [min A of G | G :!=: 1 & B \subset 'N(G)].
+
+Definition simple A := minnormal A A.
+
+Definition chief_factor A V U := maxnormal V U A && (U <| A).
+End GroupDefs.
+
+Arguments Scope subnormal [_ group_scope group_scope].
+Arguments Scope invariant_factor [_ group_scope group_scope group_scope].
+Arguments Scope stable_factor [_ group_scope group_scope group_scope].
+Arguments Scope central_factor [_ group_scope group_scope group_scope].
+Arguments Scope maximal [_ group_scope group_scope].
+Arguments Scope maximal_eq [_ group_scope group_scope].
+Arguments Scope maxnormal [_ group_scope group_scope group_scope].
+Arguments Scope minnormal [_ group_scope group_scope].
+Arguments Scope simple [_ group_scope].
+Arguments Scope chief_factor [_ group_scope group_scope group_scope].
+Prenex Implicits subnormal maximal simple.
+
+Notation "H <|<| G" := (subnormal H G)
+ (at level 70, no associativity) : group_scope.
+
+Notation "A .-invariant" := (invariant_factor A)
+ (at level 2, format "A .-invariant") : group_rel_scope.
+Notation "A .-stable" := (stable_factor A)
+ (at level 2, format "A .-stable") : group_rel_scope.
+Notation "A .-central" := (central_factor A)
+ (at level 2, format "A .-central") : group_rel_scope.
+Notation "G .-chief" := (chief_factor G)
+ (at level 2, format "G .-chief") : group_rel_scope.
+
+Arguments Scope group_rel_of [_ group_rel_scope Group_scope Group_scope].
+
+Notation "r .-series" := (path (rel_of_simpl_rel (group_rel_of r)))
+ (at level 2, format "r .-series") : group_scope.
+
+Section Subnormal.
+
+Variable gT : finGroupType.
+Implicit Types (A B C D : {set gT}) (G H K : {group gT}).
+
+Let setIgr H G := (G :&: H)%G.
+Let sub_setIgr G H : G \subset H -> G = setIgr H G.
+Proof. by move/setIidPl/group_inj. Qed.
+
+Let path_setIgr H G s :
+ normal.-series H s -> normal.-series (setIgr G H) (map (setIgr G) s).
+Proof.
+elim: s H => //= K s IHs H /andP[/andP[sHK nHK] Ksn].
+by rewrite /normal setSI ?normsIG ?IHs.
+Qed.
+
+Lemma subnormalP H G :
+ reflect (exists2 s, normal.-series H s & last H s = G) (H <|<| G).
+Proof.
+apply: (iffP andP) => [[sHG snHG] | [s Hsn <-{G}]].
+ elim: {G}#|G| {-2}G sHG snHG => [|m IHm] G sHG.
+ by exists [::]; last by apply/eqP; rewrite eq_sym.
+ rewrite iterSr => /IHm[|s Hsn defG].
+ by rewrite sub_gen // class_supportEr (bigD1 1) //= conjsg1 subsetUl.
+ exists (rcons s G); rewrite ?last_rcons // -cats1 cat_path Hsn defG /=.
+ rewrite /normal gen_subG class_support_subG //=.
+ by rewrite norms_gen ?class_support_norm.
+set f := fun _ => <<_>>; have idf: iter _ f H == H.
+ by elim=> //= m IHm; rewrite (eqP IHm) /f class_support_id genGid.
+elim: {s}(size s) {-2}s (eqxx (size s)) Hsn => [[] //= | m IHm s].
+case/lastP: s => // s G; rewrite size_rcons last_rcons -cats1 cat_path /=.
+set K := last H s => def_m /and3P[Hsn /andP[sKG nKG] _].
+have:= sKG; rewrite subEproper; case/predU1P=> [<-|prKG]; first exact: IHm.
+pose L := [group of f G].
+have sHK: H \subset K by case/IHm: Hsn.
+have sLK: L \subset K by rewrite gen_subG class_support_sub_norm.
+rewrite -(subnK (proper_card (sub_proper_trans sLK prKG))) iter_add iterSr.
+have defH: H = setIgr L H by rewrite -sub_setIgr ?sub_gen ?sub_class_support.
+have: normal.-series H (map (setIgr L) s) by rewrite defH path_setIgr.
+case/IHm=> [|_]; first by rewrite size_map.
+by rewrite {1 2}defH last_map (subset_trans sHK) //= (setIidPr sLK) => /eqP->.
+Qed.
+
+Lemma subnormal_refl G : G <|<| G.
+Proof. by apply/subnormalP; exists [::]. Qed.
+
+Lemma subnormal_trans K H G : H <|<| K -> K <|<| G -> H <|<| G.
+Proof.
+case/subnormalP=> [s1 Hs1 <-] /subnormalP[s2 Hs12 <-].
+by apply/subnormalP; exists (s1 ++ s2); rewrite ?last_cat // cat_path Hs1.
+Qed.
+
+Lemma normal_subnormal H G : H <| G -> H <|<| G.
+Proof. by move=> nsHG; apply/subnormalP; exists [:: G]; rewrite //= nsHG. Qed.
+
+Lemma setI_subnormal G H K : K \subset G -> H <|<| G -> H :&: K <|<| K.
+Proof.
+move=> sKG /subnormalP[s Hs defG]; apply/subnormalP.
+exists (map (setIgr K) s); first exact: path_setIgr.
+rewrite (last_map (setIgr K)) defG.
+by apply: val_inj; rewrite /= (setIidPr sKG).
+Qed.
+
+Lemma subnormal_sub G H : H <|<| G -> H \subset G.
+Proof. by case/andP. Qed.
+
+Lemma invariant_subnormal A G H :
+ A \subset 'N(G) -> A \subset 'N(H) -> H <|<| G ->
+ exists2 s, (A.-invariant).-series H s & last H s = G.
+Proof.
+move=> nGA nHA /andP[]; move: #|G| => m.
+elim: m => [|m IHm] in G nGA * => sHG.
+ by rewrite eq_sym; exists [::]; last exact/eqP.
+rewrite iterSr; set K := <<_>>.
+have nKA: A \subset 'N(K) by rewrite norms_gen ?norms_class_support.
+have sHK: H \subset K by rewrite sub_gen ?sub_class_support.
+case/IHm=> // s Hsn defK; exists (rcons s G); last by rewrite last_rcons.
+rewrite rcons_path Hsn !andbA defK nGA nKA /= -/K.
+by rewrite gen_subG class_support_subG ?norms_gen ?class_support_norm.
+Qed.
+
+Lemma subnormalEsupport G H :
+ H <|<| G -> H :=: G \/ <<class_support H G>> \proper G.
+Proof.
+case/andP=> sHG; set K := <<_>> => /eqP <-.
+have: K \subset G by rewrite gen_subG class_support_subG.
+rewrite subEproper; case/predU1P=> [defK|]; [left | by right].
+by elim: #|G| => //= _ ->.
+Qed.
+
+Lemma subnormalEr G H : H <|<| G ->
+ H :=: G \/ (exists K : {group gT}, [/\ H <|<| K, K <| G & K \proper G]).
+Proof.
+case/subnormalP=> s Hs <-{G}.
+elim/last_ind: s Hs => [|s G IHs]; first by left.
+rewrite last_rcons -cats1 cat_path /= andbT; set K := last H s.
+case/andP=> Hs nsKG; have:= normal_sub nsKG; rewrite subEproper.
+case/predU1P=> [<- | prKG]; [exact: IHs | right; exists K; split=> //].
+by apply/subnormalP; exists s.
+Qed.
+
+Lemma subnormalEl G H : H <|<| G ->
+ H :=: G \/ (exists K : {group gT}, [/\ H <| K, K <|<| G & H \proper K]).
+Proof.
+case/subnormalP=> s Hs <-{G}; elim: s H Hs => /= [|K s IHs] H; first by left.
+case/andP=> nsHK Ks; have:= normal_sub nsHK; rewrite subEproper.
+case/predU1P=> [-> | prHK]; [exact: IHs | right; exists K; split=> //].
+by apply/subnormalP; exists s.
+Qed.
+
+End Subnormal.
+
+Implicit Arguments subnormalP [gT G H].
+Prenex Implicits subnormalP.
+
+Section MorphSubNormal.
+
+Variable gT : finGroupType.
+Implicit Type G H K : {group gT}.
+
+Lemma morphim_subnormal (rT : finGroupType) G (f : {morphism G >-> rT}) H K :
+ H <|<| K -> f @* H <|<| f @* K.
+Proof.
+case/subnormalP => s Hs <-{K}; apply/subnormalP.
+elim: s H Hs => [|K s IHs] H /=; first by exists [::].
+case/andP=> nsHK /IHs[fs Hfs <-].
+by exists ([group of f @* K] :: fs); rewrite /= ?morphim_normal.
+Qed.
+
+Lemma quotient_subnormal H G K : G <|<| K -> G / H <|<| K / H.
+Proof. exact: morphim_subnormal. Qed.
+
+End MorphSubNormal.
+
+Section MaxProps.
+
+Variable gT : finGroupType.
+Implicit Types G H M : {group gT}.
+
+Lemma maximal_eqP M G :
+ reflect (M \subset G /\
+ forall H, M \subset H -> H \subset G -> H :=: M \/ H :=: G)
+ (maximal_eq M G).
+Proof.
+rewrite subEproper /maximal_eq; case: eqP => [->|_]; first left.
+ by split=> // H sGH sHG; right; apply/eqP; rewrite eqEsubset sHG.
+apply: (iffP maxgroupP) => [] [sMG maxM]; split=> // H.
+ by move/maxM=> maxMH; rewrite subEproper; case/predU1P; auto.
+by rewrite properEneq => /andP[/eqP neHG sHG] /maxM[].
+Qed.
+
+Lemma maximal_exists H G :
+ H \subset G ->
+ H :=: G \/ (exists2 M : {group gT}, maximal M G & H \subset M).
+Proof.
+rewrite subEproper; case/predU1P=> sHG; first by left.
+suff [M *]: {M : {group gT} | maximal M G & H \subset M} by right; exists M.
+exact: maxgroup_exists.
+Qed.
+
+Lemma mulg_normal_maximal G M H :
+ M <| G -> maximal M G -> H \subset G -> ~~ (H \subset M) -> (M * H = G)%g.
+Proof.
+case/andP=> sMG nMG /maxgroupP[_ maxM] sHG not_sHM.
+apply/eqP; rewrite eqEproper mul_subG // -norm_joinEr ?(subset_trans sHG) //.
+by apply: contra not_sHM => /maxM <-; rewrite ?joing_subl ?joing_subr.
+Qed.
+
+End MaxProps.
+
+Section MinProps.
+
+Variable gT : finGroupType.
+Implicit Types G H M : {group gT}.
+
+Lemma minnormal_exists G H : H :!=: 1 -> G \subset 'N(H) ->
+ {M : {group gT} | minnormal M G & M \subset H}.
+Proof. by move=> ntH nHG; apply: mingroup_exists (H) _; rewrite ntH. Qed.
+
+End MinProps.
+
+Section MorphPreMax.
+
+Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+Variables (M G : {group rT}).
+Hypotheses (dM : M \subset f @* D) (dG : G \subset f @* D).
+
+Lemma morphpre_maximal : maximal (f @*^-1 M) (f @*^-1 G) = maximal M G.
+Proof.
+apply/maxgroupP/maxgroupP; rewrite morphpre_proper //= => [] [ltMG maxM].
+ split=> // H ltHG sMH; have dH := subset_trans (proper_sub ltHG) dG.
+ rewrite -(morphpreK dH) [f @*^-1 H]maxM ?morphpreK ?morphpreSK //.
+ by rewrite morphpre_proper.
+split=> // H ltHG sMH.
+have dH: H \subset D := subset_trans (proper_sub ltHG) (subsetIl D _).
+have defH: f @*^-1 (f @* H) = H.
+ by apply: morphimGK dH; apply: subset_trans sMH; exact: ker_sub_pre.
+rewrite -defH morphpre_proper ?morphimS // in ltHG.
+by rewrite -defH [f @* H]maxM // -(morphpreK dM) morphimS.
+Qed.
+
+Lemma morphpre_maximal_eq : maximal_eq (f @*^-1 M) (f @*^-1 G) = maximal_eq M G.
+Proof. by rewrite /maximal_eq morphpre_maximal !eqEsubset !morphpreSK. Qed.
+
+End MorphPreMax.
+
+Section InjmMax.
+
+Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+Variables M G L : {group gT}.
+
+Hypothesis injf : 'injm f.
+Hypotheses (dM : M \subset D) (dG : G \subset D) (dL : L \subset D).
+
+Lemma injm_maximal : maximal (f @* M) (f @* G) = maximal M G.
+Proof.
+rewrite -(morphpre_invm injf) -(morphpre_invm injf G).
+by rewrite morphpre_maximal ?morphim_invm.
+Qed.
+
+Lemma injm_maximal_eq : maximal_eq (f @* M) (f @* G) = maximal_eq M G.
+Proof. by rewrite /maximal_eq injm_maximal // injm_eq. Qed.
+
+Lemma injm_maxnormal : maxnormal (f @* M) (f @* G) (f @* L) = maxnormal M G L.
+Proof.
+pose injfm := (injm_proper injf, injm_norms, injmSK injf, subsetIl).
+apply/maxgroupP/maxgroupP; rewrite !injfm // => [[nML maxM]].
+ split=> // H nHL sMH; have [/proper_sub sHG _] := andP nHL.
+ have dH := subset_trans sHG dG; apply: (injm_morphim_inj injf) => //.
+ by apply: maxM; rewrite !injfm.
+split=> // fH nHL sMH; have [/proper_sub sfHG _] := andP nHL.
+have{sfHG} dfH: fH \subset f @* D := subset_trans sfHG (morphim_sub f G).
+by rewrite -(morphpreK dfH) !injfm // in nHL sMH *; rewrite (maxM _ nHL).
+Qed.
+
+Lemma injm_minnormal : minnormal (f @* M) (f @* G) = minnormal M G.
+Proof.
+pose injfm := (morphim_injm_eq1 injf, injm_norms, injmSK injf, subsetIl).
+apply/mingroupP/mingroupP; rewrite !injfm // => [[nML minM]].
+ split=> // H nHG sHM; have dH := subset_trans sHM dM.
+ by apply: (injm_morphim_inj injf) => //; apply: minM; rewrite !injfm.
+split=> // fH nHG sHM; have dfH := subset_trans sHM (morphim_sub f M).
+by rewrite -(morphpreK dfH) !injfm // in nHG sHM *; rewrite (minM _ nHG).
+Qed.
+
+End InjmMax.
+
+Section QuoMax.
+
+Variables (gT : finGroupType) (K G H : {group gT}).
+
+Lemma cosetpre_maximal (Q R : {group coset_of K}) :
+ maximal (coset K @*^-1 Q) (coset K @*^-1 R) = maximal Q R.
+Proof. by rewrite morphpre_maximal ?sub_im_coset. Qed.
+
+Lemma cosetpre_maximal_eq (Q R : {group coset_of K}) :
+ maximal_eq (coset K @*^-1 Q) (coset K @*^-1 R) = maximal_eq Q R.
+Proof. by rewrite /maximal_eq !eqEsubset !cosetpreSK cosetpre_maximal. Qed.
+
+Lemma quotient_maximal :
+ K <| G -> K <| H -> maximal (G / K) (H / K) = maximal G H.
+Proof. by move=> nKG nKH; rewrite -cosetpre_maximal ?quotientGK. Qed.
+
+Lemma quotient_maximal_eq :
+ K <| G -> K <| H -> maximal_eq (G / K) (H / K) = maximal_eq G H.
+Proof. by move=> nKG nKH; rewrite -cosetpre_maximal_eq ?quotientGK. Qed.
+
+Lemma maximalJ x : maximal (G :^ x) (H :^ x) = maximal G H.
+Proof.
+rewrite -{1}(setTI G) -{1}(setTI H) -!morphim_conj.
+by rewrite injm_maximal ?subsetT ?injm_conj.
+Qed.
+
+Lemma maximal_eqJ x : maximal_eq (G :^ x) (H :^ x) = maximal_eq G H.
+Proof. by rewrite /maximal_eq !eqEsubset !conjSg maximalJ. Qed.
+
+End QuoMax.
+
+Section MaxNormalProps.
+
+Variables (gT : finGroupType).
+Implicit Types (A B C : {set gT}) (G H K L M : {group gT}).
+
+Lemma maxnormal_normal A B : maxnormal A B B -> A <| B.
+Proof.
+by case/maxsetP=> /and3P[/gen_set_id /= -> pAB nAB]; rewrite /normal proper_sub.
+Qed.
+
+Lemma maxnormal_proper A B C : maxnormal A B C -> A \proper B.
+Proof.
+by case/maxsetP=> /and3P[gA pAB _] _; exact: (sub_proper_trans (subset_gen A)).
+Qed.
+
+Lemma maxnormal_sub A B C : maxnormal A B C -> A \subset B.
+Proof.
+by move=> maxA; rewrite proper_sub //; exact: (maxnormal_proper maxA).
+Qed.
+
+Lemma ex_maxnormal_ntrivg G : G :!=: 1-> {N : {group gT} | maxnormal N G G}.
+Proof.
+move=> ntG; apply: ex_maxgroup; exists [1 gT]%G; rewrite norm1 proper1G.
+by rewrite subsetT ntG.
+Qed.
+
+Lemma maxnormalM G H K :
+ maxnormal H G G -> maxnormal K G G -> H :<>: K -> H * K = G.
+Proof.
+move=> maxH maxK /eqP; apply: contraNeq => ltHK_G.
+have [nsHG nsKG] := (maxnormal_normal maxH, maxnormal_normal maxK).
+have cHK: commute H K.
+ exact: normC (subset_trans (normal_sub nsHG) (normal_norm nsKG)).
+wlog suffices: H K {maxH} maxK nsHG nsKG cHK ltHK_G / H \subset K.
+ by move=> IH; rewrite eqEsubset !IH // -cHK.
+have{maxK} /maxgroupP[_ maxK] := maxK.
+apply/joing_idPr/maxK; rewrite ?joing_subr //= comm_joingE //.
+by rewrite properEneq ltHK_G; exact: normalM.
+Qed.
+
+Lemma maxnormal_minnormal G L M :
+ G \subset 'N(M) -> L \subset 'N(G) -> maxnormal M G L ->
+ minnormal (G / M) (L / M).
+Proof.
+move=> nMG nGL /maxgroupP[/andP[/andP[sMG ltMG] nML] maxM]; apply/mingroupP.
+rewrite -subG1 quotient_sub1 ?ltMG ?quotient_norms //.
+split=> // Hb /andP[ntHb nHbL]; have nsMG: M <| G by exact/andP.
+case/inv_quotientS=> // H defHb sMH sHG; rewrite defHb; congr (_ / M).
+apply/eqP; rewrite eqEproper sHG /=; apply: contra ntHb => ltHG.
+have nsMH: M <| H := normalS sMH sHG nsMG.
+rewrite defHb quotientS1 // (maxM H) // ltHG /= -(quotientGK nsMH) -defHb.
+exact: norm_quotient_pre.
+Qed.
+
+Lemma minnormal_maxnormal G L M :
+ M <| G -> L \subset 'N(M) -> minnormal (G / M) (L / M) -> maxnormal M G L.
+Proof.
+case/andP=> sMG nMG nML /mingroupP[/andP[/= ntGM _] minGM]; apply/maxgroupP.
+split=> [|H /andP[/andP[sHG ltHG] nHL] sMH].
+ by rewrite /proper sMG nML andbT; apply: contra ntGM => /quotientS1 ->.
+apply/eqP; rewrite eqEsubset sMH andbT -quotient_sub1 ?(subset_trans sHG) //.
+rewrite subG1; apply: contraR ltHG => ntHM; rewrite -(quotientSGK nMG) //.
+by rewrite (minGM (H / M)%G) ?quotientS // ntHM quotient_norms.
+Qed.
+
+End MaxNormalProps.
+
+Section Simple.
+
+Implicit Types gT rT : finGroupType.
+
+Lemma simpleP gT (G : {group gT}) :
+ reflect (G :!=: 1 /\ forall H : {group gT}, H <| G -> H :=: 1 \/ H :=: G)
+ (simple G).
+Proof.
+apply: (iffP mingroupP); rewrite normG andbT => [[ntG simG]].
+ split=> // N /andP[sNG nNG].
+ by case: (eqsVneq N 1) => [|ntN]; [left | right; apply: simG; rewrite ?ntN].
+split=> // N /andP[ntN nNG] sNG.
+by case: (simG N) ntN => // [|->]; [exact/andP | case/eqP].
+Qed.
+
+Lemma quotient_simple gT (G H : {group gT}) :
+ H <| G -> simple (G / H) = maxnormal H G G.
+Proof.
+move=> nsHG; have nGH := normal_norm nsHG.
+by apply/idP/idP; [exact: minnormal_maxnormal | exact: maxnormal_minnormal].
+Qed.
+
+Lemma isog_simple gT rT (G : {group gT}) (M : {group rT}) :
+ G \isog M -> simple G = simple M.
+Proof.
+move=> eqGM; wlog suffices: gT rT G M eqGM / simple M -> simple G.
+ by move=> IH; apply/idP/idP; apply: IH; rewrite // isog_sym.
+case/isogP: eqGM => f injf <- /simpleP[ntGf simGf].
+apply/simpleP; split=> [|N nsNG]; first by rewrite -(morphim_injm_eq1 injf).
+rewrite -(morphim_invm injf (normal_sub nsNG)).
+have: f @* N <| f @* G by rewrite morphim_normal.
+by case/simGf=> /= ->; [left | right]; rewrite (morphim1, morphim_invm).
+Qed.
+
+Lemma simple_maxnormal gT (G : {group gT}) : simple G = maxnormal 1 G G.
+Proof.
+by rewrite -quotient_simple ?normal1 // -(isog_simple (quotient1_isog G)).
+Qed.
+
+End Simple.
+
+Section Chiefs.
+
+Variable gT : finGroupType.
+Implicit Types G H U V : {group gT}.
+
+Lemma chief_factor_minnormal G V U :
+ chief_factor G V U -> minnormal (U / V) (G / V).
+Proof.
+case/andP=> maxV /andP[sUG nUG]; apply: maxnormal_minnormal => //.
+by have /andP[_ nVG] := maxgroupp maxV; exact: subset_trans sUG nVG.
+Qed.
+
+Lemma acts_irrQ G U V :
+ G \subset 'N(V) -> V <| U ->
+ acts_irreducibly G (U / V) 'Q = minnormal (U / V) (G / V).
+Proof.
+move=> nVG nsVU; apply/mingroupP/mingroupP; case=> /andP[->] /=.
+ rewrite astabsQ // subsetI nVG /= => nUG minUV.
+ rewrite quotient_norms //; split=> // H /andP[ntH nHG] sHU.
+ by apply: minUV (sHU); rewrite ntH -(cosetpreK H) actsQ // norm_quotient_pre.
+rewrite sub_quotient_pre // => nUG minU; rewrite astabsQ //.
+rewrite (subset_trans nUG); last first.
+ by rewrite subsetI subsetIl /= -{2}(quotientGK nsVU) morphpre_norm.
+split=> // H /andP[ntH nHG] sHU.
+rewrite -{1}(cosetpreK H) astabsQ ?normal_cosetpre ?subsetI ?nVG //= in nHG.
+apply: minU sHU; rewrite ntH; apply: subset_trans (quotientS _ nHG) _.
+by rewrite -{2}(cosetpreK H) quotient_norm.
+Qed.
+
+Lemma chief_series_exists H G :
+ H <| G -> {s | (G.-chief).-series 1%G s & last 1%G s = H}.
+Proof.
+elim: {H}_.+1 {-2}H (ltnSn #|H|) => // m IHm U leUm nsUG.
+have [-> | ntU] := eqVneq U 1%G; first by exists [::].
+have [V maxV]: {V : {group gT} | maxnormal V U G}.
+ by apply: ex_maxgroup; exists 1%G; rewrite proper1G ntU norms1.
+have /andP[ltVU nVG] := maxgroupp maxV.
+have [||s ch_s defV] := IHm V; first exact: leq_trans (proper_card ltVU) _.
+ by rewrite /normal (subset_trans (proper_sub ltVU) (normal_sub nsUG)).
+exists (rcons s U); last by rewrite last_rcons.
+by rewrite rcons_path defV /= ch_s /chief_factor; exact/and3P.
+Qed.
+
+End Chiefs.
+
+Section Central.
+
+Variables (gT : finGroupType) (G : {group gT}).
+Implicit Types H K : {group gT}.
+
+Lemma central_factor_central H K :
+ central_factor G H K -> (K / H) \subset 'Z(G / H).
+Proof. by case/and3P=> /quotient_cents2r *; rewrite subsetI quotientS. Qed.
+
+
+Lemma central_central_factor H K :
+ (K / H) \subset 'Z(G / H) -> H <| K -> H <| G -> central_factor G H K.
+Proof.
+case/subsetIP=> sKGb cGKb /andP[sHK nHK] /andP[sHG nHG].
+by rewrite /central_factor -quotient_cents2 // cGKb sHK -(quotientSGK nHK).
+Qed.
+
+End Central.
diff --git a/mathcomp/solvable/hall.v b/mathcomp/solvable/hall.v
new file mode 100644
index 0000000..e097765
--- /dev/null
+++ b/mathcomp/solvable/hall.v
@@ -0,0 +1,895 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype finset.
+Require Import prime fingroup morphism automorphism quotient action gproduct.
+Require Import commutator center pgroup finmodule nilpotent sylow.
+Require Import abelian maximal.
+
+(*****************************************************************************)
+(* In this files we prove the Schur-Zassenhaus splitting and transitivity *)
+(* theorems (under solvability assumptions), then derive P. Hall's *)
+(* generalization of Sylow's theorem to solvable groups and its corollaries, *)
+(* in particular the theory of coprime action. We develop both the theory of *)
+(* coprime action of a solvable group on Sylow subgroups (as in Aschbacher *)
+(* 18.7), and that of coprime action on Hall subgroups of a solvable group *)
+(* as per B & G, Proposition 1.5; however we only support external group *)
+(* action (as opposed to internal action by conjugation) for the latter case *)
+(* because it is much harder to apply in practice. *)
+(*****************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Hall.
+
+Implicit Type gT : finGroupType.
+
+Theorem SchurZassenhaus_split gT (G H : {group gT}) :
+ Hall G H -> H <| G -> [splits G, over H].
+Proof.
+move: {2}_.+1 (ltnSn #|G|) => n; elim: n => // n IHn in gT G H *.
+rewrite ltnS => Gn hallH nsHG; have [sHG nHG] := andP nsHG.
+have [-> | [p pr_p pH]] := trivgVpdiv H.
+ by apply/splitsP; exists G; rewrite inE -subG1 subsetIl mul1g eqxx.
+have [P sylP] := Sylow_exists p H.
+case nPG: (P <| G); last first.
+ pose N := ('N_G(P))%G; have sNG: N \subset G by rewrite subsetIl.
+ have eqHN_G: H * N = G by exact: Frattini_arg sylP.
+ pose H' := (H :&: N)%G.
+ have nsH'N: H' <| N.
+ by rewrite /normal subsetIr normsI ?normG ?(subset_trans sNG).
+ have eq_iH: #|G : H| = #|N| %/ #|H'|.
+ rewrite -divgS // -(divnMl (cardG_gt0 H')) mulnC -eqHN_G.
+ by rewrite -mul_cardG (mulnC #|H'|) divnMl // cardG_gt0.
+ have hallH': Hall N H'.
+ rewrite /Hall -divgS subsetIr //= -eq_iH.
+ by case/andP: hallH => _; apply: coprimeSg; exact: subsetIl.
+ have: [splits N, over H'].
+ apply: IHn hallH' nsH'N; apply: {n}leq_trans Gn.
+ rewrite proper_card // properEneq sNG andbT; apply/eqP=> eqNG.
+ by rewrite -eqNG normal_subnorm (subset_trans (pHall_sub sylP)) in nPG.
+ case/splitsP=> K /complP[tiKN eqH'K].
+ have sKN: K \subset N by rewrite -(mul1g K) -eqH'K mulSg ?sub1set.
+ apply/splitsP; exists K; rewrite inE -subG1; apply/andP; split.
+ by rewrite /= -(setIidPr sKN) setIA tiKN.
+ by rewrite eqEsubset -eqHN_G mulgS // -eqH'K mulGS mulSg ?subsetIl.
+pose Z := 'Z(P); pose Gbar := G / Z; pose Hbar := H / Z.
+have sZP: Z \subset P by exact: center_sub.
+have sZH: Z \subset H by exact: subset_trans (pHall_sub sylP).
+have sZG: Z \subset G by exact: subset_trans sHG.
+have nZG: Z <| G by apply: char_normal_trans nPG; exact: center_char.
+have nZH: Z <| H by exact: normalS nZG.
+have nHGbar: Hbar <| Gbar by exact: morphim_normal.
+have hallHbar: Hall Gbar Hbar by apply: morphim_Hall (normal_norm _) _.
+have: [splits Gbar, over Hbar].
+ apply: IHn => //; apply: {n}leq_trans Gn; rewrite ltn_quotient //.
+ apply/eqP=> /(trivg_center_pgroup (pHall_pgroup sylP))/eqP.
+ rewrite trivg_card1 (card_Hall sylP) p_part -(expn0 p).
+ by rewrite eqn_exp2l ?prime_gt1 // lognE pH pr_p cardG_gt0.
+case/splitsP=> Kbar /complP[tiHKbar eqHKbar].
+have: Kbar \subset Gbar by rewrite -eqHKbar mulG_subr.
+case/inv_quotientS=> //= ZK quoZK sZZK sZKG.
+have nZZK: Z <| ZK by exact: normalS nZG.
+have cardZK: #|ZK| = (#|Z| * #|G : H|)%N.
+ rewrite -(Lagrange sZZK); congr (_ * _)%N.
+ rewrite -card_quotient -?quoZK; last by case/andP: nZZK.
+ rewrite -(divgS sHG) -(Lagrange sZG) -(Lagrange sZH) divnMl //.
+ rewrite -!card_quotient ?normal_norm //= -/Gbar -/Hbar.
+ by rewrite -eqHKbar (TI_cardMg tiHKbar) mulKn.
+have: [splits ZK, over Z].
+ rewrite (Gaschutz_split nZZK _ sZZK) ?center_abelian //; last first.
+ rewrite -divgS // cardZK mulKn ?cardG_gt0 //.
+ by case/andP: hallH => _; exact: coprimeSg.
+ by apply/splitsP; exists 1%G; rewrite inE -subG1 subsetIr mulg1 eqxx.
+case/splitsP=> K /complP[tiZK eqZK].
+have sKZK: K \subset ZK by rewrite -(mul1g K) -eqZK mulSg ?sub1G.
+have tiHK: H :&: K = 1.
+ apply/trivgP; rewrite /= -(setIidPr sKZK) setIA -tiZK setSI //.
+ rewrite -quotient_sub1; last by rewrite subIset 1?normal_norm.
+ by rewrite /= quotientGI //= -quoZK tiHKbar.
+apply/splitsP; exists K; rewrite inE tiHK ?eqEcard subxx leqnn /=.
+rewrite mul_subG ?(subset_trans sKZK) //= TI_cardMg //.
+rewrite -(@mulKn #|K| #|Z|) ?cardG_gt0 // -TI_cardMg // eqZK.
+by rewrite cardZK mulKn ?cardG_gt0 // Lagrange.
+Qed.
+
+Theorem SchurZassenhaus_trans_sol gT (H K K1 : {group gT}) :
+ solvable H -> K \subset 'N(H) -> K1 \subset H * K ->
+ coprime #|H| #|K| -> #|K1| = #|K| ->
+ exists2 x, x \in H & K1 :=: K :^ x.
+Proof.
+move: {2}_.+1 (ltnSn #|H|) => n; elim: n => // n IHn in gT H K K1 *.
+rewrite ltnS => leHn solH nHK; have [-> | ] := eqsVneq H 1.
+ rewrite mul1g => sK1K _ eqK1K; exists 1; first exact: set11.
+ by apply/eqP; rewrite conjsg1 eqEcard sK1K eqK1K /=.
+pose G := (H <*> K)%G.
+have defG: G :=: H * K by rewrite -normC // -norm_joinEl // joingC.
+have sHG: H \subset G by exact: joing_subl.
+have sKG: K \subset G by exact: joing_subr.
+have nsHG: H <| G by rewrite /(H <| G) sHG join_subG normG.
+case/(solvable_norm_abelem solH nsHG)=> M [sMH nsMG ntM] /and3P[_ abelM _].
+have [sMG nMG] := andP nsMG; rewrite -defG => sK1G coHK oK1K.
+have nMsG (L : {set gT}): L \subset G -> L \subset 'N(M).
+ by move/subset_trans->.
+have [coKM coHMK]: coprime #|M| #|K| /\ coprime #|H / M| #|K|.
+ by apply/andP; rewrite -coprime_mull card_quotient ?nMsG ?Lagrange.
+have oKM (K' : {group gT}): K' \subset G -> #|K'| = #|K| -> #|K' / M| = #|K|.
+ move=> sK'G oK'.
+ rewrite -quotientMidr -?norm_joinEl ?card_quotient ?nMsG //; last first.
+ by rewrite gen_subG subUset sK'G.
+ rewrite -divgS /=; last by rewrite -gen_subG genS ?subsetUr.
+ by rewrite norm_joinEl ?nMsG // coprime_cardMg ?mulnK // oK' coprime_sym.
+have [xb]: exists2 xb, xb \in H / M & K1 / M = (K / M) :^ xb.
+ apply: IHn; try by rewrite (quotient_sol, morphim_norms, oKM K) ?(oKM K1).
+ by apply: leq_trans leHn; rewrite ltn_quotient.
+ by rewrite -morphimMl ?nMsG // -defG morphimS.
+case/morphimP=> x nMx Hx ->{xb} eqK1Kx; pose K2 := (K :^ x)%G.
+have{eqK1Kx} eqK12: K1 / M = K2 / M by rewrite quotientJ.
+suff [y My ->]: exists2 y, y \in M & K1 :=: K2 :^ y.
+ by exists (x * y); [rewrite groupMl // (subsetP sMH) | rewrite conjsgM].
+have nMK1: K1 \subset 'N(M) by exact: nMsG.
+have defMK: M * K1 = M <*> K1 by rewrite -normC // -norm_joinEl // joingC.
+have sMKM: M \subset M <*> K1 by rewrite joing_subl.
+have nMKM: M <| M <*> K1 by rewrite normalYl.
+have trMK1: M :&: K1 = 1 by rewrite coprime_TIg ?oK1K.
+have trMK2: M :&: K2 = 1 by rewrite coprime_TIg ?cardJg ?oK1K.
+apply: (Gaschutz_transitive nMKM _ sMKM) => //=; last 2 first.
+- by rewrite inE trMK1 defMK !eqxx.
+- by rewrite -!(setIC M) trMK1.
+- by rewrite -divgS //= -defMK coprime_cardMg oK1K // mulKn.
+rewrite inE trMK2 eqxx eq_sym eqEcard /= -defMK andbC.
+by rewrite !coprime_cardMg ?cardJg ?oK1K ?leqnn //= mulGS -quotientSK -?eqK12.
+Qed.
+
+Lemma SchurZassenhaus_trans_actsol gT (G A B : {group gT}) :
+ solvable A -> A \subset 'N(G) -> B \subset A <*> G ->
+ coprime #|G| #|A| -> #|A| = #|B| ->
+ exists2 x, x \in G & B :=: A :^ x.
+Proof.
+set AG := A <*> G; move: {2}_.+1 (ltnSn #|AG|) => n.
+elim: n => // n IHn in gT A B G AG *.
+rewrite ltnS => leAn solA nGA sB_AG coGA oAB.
+have [A1 | ntA] := eqsVneq A 1.
+ by exists 1; rewrite // conjsg1 A1 (@card1_trivg _ B) // -oAB A1 cards1.
+have [M [sMA nsMA ntM]] := solvable_norm_abelem solA (normal_refl A) ntA.
+case/is_abelemP=> q q_pr /abelem_pgroup qM; have nMA := normal_norm nsMA.
+have defAG: AG = A * G := norm_joinEl nGA.
+have sA_AG: A \subset AG := joing_subl _ _.
+have sG_AG: G \subset AG := joing_subr _ _.
+have sM_AG := subset_trans sMA sA_AG.
+have oAG: #|AG| = (#|A| * #|G|)%N by rewrite defAG coprime_cardMg 1?coprime_sym.
+have q'G: #|G|`_q = 1%N.
+ rewrite part_p'nat ?p'natE -?prime_coprime // coprime_sym.
+ have [_ _ [k oM]] := pgroup_pdiv qM ntM.
+ by rewrite -(@coprime_pexpr k.+1) // -oM (coprimegS sMA).
+have coBG: coprime #|B| #|G| by rewrite -oAB coprime_sym.
+have defBG: B * G = AG.
+ by apply/eqP; rewrite eqEcard mul_subG ?sG_AG //= oAG oAB coprime_cardMg.
+case nMG: (G \subset 'N(M)).
+ have nsM_AG: M <| AG by rewrite /normal sM_AG join_subG nMA.
+ have nMB: B \subset 'N(M) := subset_trans sB_AG (normal_norm nsM_AG).
+ have sMB: M \subset B.
+ have [Q sylQ]:= Sylow_exists q B; have sQB := pHall_sub sylQ.
+ apply: subset_trans (normal_sub_max_pgroup (Hall_max _) qM nsM_AG) (sQB).
+ rewrite pHallE (subset_trans sQB) //= oAG partnM // q'G muln1 oAB.
+ by rewrite (card_Hall sylQ).
+ have defAGq: AG / M = (A / M) <*> (G / M).
+ by rewrite quotient_gen ?quotientU ?subUset ?nMA.
+ have: B / M \subset (A / M) <*> (G / M) by rewrite -defAGq quotientS.
+ case/IHn; rewrite ?morphim_sol ?quotient_norms ?coprime_morph //.
+ - by rewrite -defAGq (leq_trans _ leAn) ?ltn_quotient.
+ - by rewrite !card_quotient // -!divgS // oAB.
+ move=> Mx; case/morphimP=> x Nx Gx ->{Mx} //; rewrite -quotientJ //= => defBq.
+ exists x => //; apply: quotient_inj defBq; first by rewrite /normal sMB.
+ by rewrite -(normsP nMG x Gx) /normal normJ !conjSg.
+pose K := M <*> G; pose R := K :&: B; pose N := 'N_G(M).
+have defK: K = M * G by rewrite -norm_joinEl ?(subset_trans sMA).
+have oK: #|K| = (#|M| * #|G|)%N.
+ by rewrite defK coprime_cardMg // coprime_sym (coprimegS sMA).
+have sylM: q.-Sylow(K) M.
+ by rewrite pHallE joing_subl /= oK partnM // q'G muln1 part_pnat_id.
+have sylR: q.-Sylow(K) R.
+ rewrite pHallE subsetIl /= -(card_Hall sylM) -(@eqn_pmul2r #|G|) // -oK.
+ rewrite -coprime_cardMg ?(coprimeSg _ coBG) ?subsetIr //=.
+ by rewrite group_modr ?joing_subr ?(setIidPl _) // defBG join_subG sM_AG.
+have [mx] := Sylow_trans sylM sylR.
+rewrite /= -/K defK; case/imset2P=> m x Mm Gx ->{mx}.
+rewrite conjsgM conjGid {m Mm}// => defR.
+have sNG: N \subset G := subsetIl _ _.
+have pNG: N \proper G by rewrite /proper sNG subsetI subxx nMG.
+have nNA: A \subset 'N(N) by rewrite normsI ?norms_norm.
+have: B :^ x^-1 \subset A <*> N.
+ rewrite norm_joinEl ?group_modl // -defAG subsetI !sub_conjgV -normJ -defR.
+ rewrite conjGid ?(subsetP sG_AG) // normsI ?normsG // (subset_trans sB_AG) //.
+ by rewrite join_subG normsM // -defK normsG ?joing_subr.
+do [case/IHn; rewrite ?cardJg ?(coprimeSg _ coGA) //= -/N] => [|y Ny defB].
+ rewrite joingC norm_joinEr // coprime_cardMg ?(coprimeSg sNG) //.
+ by rewrite (leq_trans _ leAn) // oAG mulnC ltn_pmul2l // proper_card.
+exists (y * x); first by rewrite groupM // (subsetP sNG).
+by rewrite conjsgM -defB conjsgKV.
+Qed.
+
+Lemma Hall_exists_subJ pi gT (G : {group gT}) :
+ solvable G -> exists2 H : {group gT}, pi.-Hall(G) H
+ & forall K : {group gT}, K \subset G -> pi.-group K ->
+ exists2 x, x \in G & K \subset H :^ x.
+Proof.
+move: {2}_.+1 (ltnSn #|G|) => n.
+elim: n gT G => // n IHn gT G; rewrite ltnS => leGn solG.
+have [-> | ntG] := eqsVneq G 1.
+ exists 1%G => [|_ /trivGP-> _]; last by exists 1; rewrite ?set11 ?sub1G.
+ by rewrite pHallE sub1G cards1 part_p'nat.
+case: (solvable_norm_abelem solG (normal_refl _)) => // M [sMG nsMG ntM].
+case/is_abelemP=> p pr_p /and3P[pM cMM _].
+pose Gb := (G / M)%G; case: (IHn _ Gb) => [||Hb]; try exact: quotient_sol.
+ by rewrite (leq_trans (ltn_quotient _ _)).
+case/and3P=> [sHbGb piHb pi'Hb'] transHb.
+case: (inv_quotientS nsMG sHbGb) => H def_H sMH sHG.
+have nMG := normal_norm nsMG; have nMH := subset_trans sHG nMG.
+have{transHb} transH (K : {group gT}):
+ K \subset G -> pi.-group K -> exists2 x, x \in G & K \subset H :^ x.
+- move=> sKG piK; have nMK := subset_trans sKG nMG.
+ case: (transHb (K / M)%G) => [||xb Gxb sKHxb]; first exact: morphimS.
+ exact: morphim_pgroup.
+ case/morphimP: Gxb => x Nx Gx /= def_x; exists x => //.
+ apply/subsetP=> y Ky.
+ have: y \in coset M y by rewrite val_coset (subsetP nMK, rcoset_refl).
+ have: coset M y \in (H :^ x) / M.
+ rewrite /quotient morphimJ //=.
+ rewrite def_x def_H in sKHxb; apply: (subsetP sKHxb); exact: mem_quotient.
+ case/morphimP=> z Nz Hxz ->.
+ rewrite val_coset //; case/rcosetP=> t Mt ->; rewrite groupMl //.
+ by rewrite mem_conjg (subsetP sMH) // -mem_conjg (normP Nx).
+have{pi'Hb'} pi'H': pi^'.-nat #|G : H|.
+ move: pi'Hb'; rewrite -!divgS // def_H !card_quotient //.
+ by rewrite -(divnMl (cardG_gt0 M)) !Lagrange.
+have [pi_p | pi'p] := boolP (p \in pi).
+ exists H => //; apply/and3P; split=> //; rewrite /pgroup.
+ by rewrite -(Lagrange sMH) -card_quotient // pnat_mul -def_H (pi_pnat pM).
+have [ltHG | leGH {n IHn leGn transH}] := ltnP #|H| #|G|.
+ case: (IHn _ H (leq_trans ltHG leGn)) => [|H1]; first exact: solvableS solG.
+ case/and3P=> sH1H piH1 pi'H1' transH1.
+ have sH1G: H1 \subset G by exact: subset_trans sHG.
+ exists H1 => [|K sKG piK].
+ apply/and3P; split => //.
+ rewrite -divgS // -(Lagrange sHG) -(Lagrange sH1H) -mulnA.
+ by rewrite mulKn // pnat_mul pi'H1'.
+ case: (transH K sKG piK) => x Gx def_K.
+ case: (transH1 (K :^ x^-1)%G) => [||y Hy def_K1].
+ - by rewrite sub_conjgV.
+ - by rewrite /pgroup cardJg.
+ exists (y * x); first by rewrite groupMr // (subsetP sHG).
+ by rewrite -(conjsgKV x K) conjsgM conjSg.
+have{leGH Gb sHbGb sHG sMH pi'H'} eqHG: H = G.
+ by apply/eqP; rewrite -val_eqE eqEcard sHG.
+have{H Hb def_H eqHG piHb nMH} hallM: pi^'.-Hall(G) M.
+ rewrite /pHall /pgroup sMG pnatNK -card_quotient //=.
+ by rewrite -eqHG -def_H (pi_pnat pM).
+case/splitsP: (SchurZassenhaus_split (pHall_Hall hallM) nsMG) => H.
+case/complP=> trMH defG.
+have sHG: H \subset G by rewrite -defG mulG_subr.
+exists H => [|K sKG piK].
+ apply: etrans hallM; rewrite /pHall sMG sHG /= -!divgS // -defG andbC.
+ by rewrite (TI_cardMg trMH) mulKn ?mulnK // pnatNK.
+pose G1 := (K <*> M)%G; pose K1 := (H :&: G1)%G.
+have nMK: K \subset 'N(M) by apply: subset_trans sKG nMG.
+have defG1: M * K = G1 by rewrite -normC -?norm_joinEl.
+have sK1G1: K1 \subset M * K by rewrite defG1 subsetIr.
+have coMK: coprime #|M| #|K|.
+ by rewrite coprime_sym (pnat_coprime piK) //; exact: (pHall_pgroup hallM).
+case: (SchurZassenhaus_trans_sol _ nMK sK1G1 coMK) => [||x Mx defK1].
+- exact: solvableS solG.
+- apply/eqP; rewrite -(eqn_pmul2l (cardG_gt0 M)) -TI_cardMg //; last first.
+ by apply/trivgP; rewrite -trMH /= setIA subsetIl.
+ rewrite -coprime_cardMg // defG1; apply/eqP; congr #|(_ : {set _})|.
+ rewrite group_modl; last by rewrite -defG1 mulG_subl.
+ by apply/setIidPr; rewrite defG gen_subG subUset sKG.
+exists x^-1; first by rewrite groupV (subsetP sMG).
+by rewrite -(_ : K1 :^ x^-1 = K) ?(conjSg, subsetIl) // defK1 conjsgK.
+Qed.
+
+End Hall.
+
+Section HallCorollaries.
+
+Variable gT : finGroupType.
+
+Corollary Hall_exists pi (G : {group gT}) :
+ solvable G -> exists H : {group gT}, pi.-Hall(G) H.
+Proof. by case/(Hall_exists_subJ pi) => H; exists H. Qed.
+
+Corollary Hall_trans pi (G H1 H2 : {group gT}) :
+ solvable G -> pi.-Hall(G) H1 -> pi.-Hall(G) H2 ->
+ exists2 x, x \in G & H1 :=: H2 :^ x.
+Proof.
+move=> solG; have [H hallH transH] := Hall_exists_subJ pi solG.
+have conjH (K : {group gT}):
+ pi.-Hall(G) K -> exists2 x, x \in G & K = (H :^ x)%G.
+- move=> hallK; have [sKG piK _] := and3P hallK.
+ case: (transH K sKG piK) => x Gx sKH; exists x => //.
+ apply/eqP; rewrite -val_eqE eqEcard sKH cardJg.
+ by rewrite (card_Hall hallH) (card_Hall hallK) /=.
+case/conjH=> x1 Gx1 ->{H1}; case/conjH=> x2 Gx2 ->{H2}.
+exists (x2^-1 * x1); first by rewrite groupMl ?groupV.
+by apply: val_inj; rewrite /= conjsgM conjsgK.
+Qed.
+
+Corollary Hall_superset pi (G K : {group gT}) :
+ solvable G -> K \subset G -> pi.-group K ->
+ exists2 H : {group gT}, pi.-Hall(G) H & K \subset H.
+Proof.
+move=> solG sKG; have [H hallH transH] := Hall_exists_subJ pi solG.
+by case/transH=> // x Gx sKHx; exists (H :^ x)%G; rewrite ?pHallJ.
+Qed.
+
+Corollary Hall_subJ pi (G H K : {group gT}) :
+ solvable G -> pi.-Hall(G) H -> K \subset G -> pi.-group K ->
+ exists2 x, x \in G & K \subset H :^ x.
+Proof.
+move=> solG HallH sKG piK; have [M HallM sKM]:= Hall_superset solG sKG piK.
+have [x Gx defM] := Hall_trans solG HallM HallH.
+by exists x; rewrite // -defM.
+Qed.
+
+Corollary Hall_Jsub pi (G H K : {group gT}) :
+ solvable G -> pi.-Hall(G) H -> K \subset G -> pi.-group K ->
+ exists2 x, x \in G & K :^ x \subset H.
+Proof.
+move=> solG HallH sKG piK; have [x Gx sKHx] := Hall_subJ solG HallH sKG piK.
+by exists x^-1; rewrite ?groupV // sub_conjgV.
+Qed.
+
+Lemma Hall_Frattini_arg pi (G K H : {group gT}) :
+ solvable K -> K <| G -> pi.-Hall(K) H -> K * 'N_G(H) = G.
+Proof.
+move=> solK /andP[sKG nKG] hallH.
+have sHG: H \subset G by apply: subset_trans sKG; case/andP: hallH.
+rewrite setIC group_modl //; apply/setIidPr/subsetP=> x Gx.
+pose H1 := (H :^ x^-1)%G.
+have hallH1: pi.-Hall(K) H1 by rewrite pHallJnorm // groupV (subsetP nKG).
+case: (Hall_trans solK hallH hallH1) => y Ky defH.
+rewrite -(mulKVg y x) mem_mulg //; apply/normP.
+by rewrite conjsgM {1}defH conjsgK conjsgKV.
+Qed.
+
+End HallCorollaries.
+
+Section InternalAction.
+
+Variables (pi : nat_pred) (gT : finGroupType).
+Implicit Types G H K A X : {group gT}.
+
+(* Part of Aschbacher (18.7.4). *)
+Lemma coprime_norm_cent A G :
+ A \subset 'N(G) -> coprime #|G| #|A| -> 'N_G(A) = 'C_G(A).
+Proof.
+move=> nGA coGA; apply/eqP; rewrite eqEsubset andbC setIS ?cent_sub //=.
+rewrite subsetI subsetIl /= (sameP commG1P trivgP) -(coprime_TIg coGA).
+rewrite subsetI commg_subr subsetIr andbT.
+move: nGA; rewrite -commg_subl; apply: subset_trans.
+by rewrite commSg ?subsetIl.
+Qed.
+
+(* This is B & G, Proposition 1.5(a) *)
+Proposition coprime_Hall_exists A G :
+ A \subset 'N(G) -> coprime #|G| #|A| -> solvable G ->
+ exists2 H : {group gT}, pi.-Hall(G) H & A \subset 'N(H).
+Proof.
+move=> nGA coGA solG; have [H hallH] := Hall_exists pi solG.
+have sG_AG: G \subset A <*> G by rewrite joing_subr.
+have nG_AG: A <*> G \subset 'N(G) by rewrite join_subG nGA normG.
+pose N := 'N_(A <*> G)(H)%G.
+have nGN: N \subset 'N(G) by rewrite subIset ?nG_AG.
+have nGN_N: G :&: N <| N by rewrite /(_ <| N) subsetIr normsI ?normG.
+have NG_AG: G * N = A <*> G.
+ by apply: Hall_Frattini_arg hallH => //; exact/andP.
+have iGN_A: #|N| %/ #|G :&: N| = #|A|.
+ rewrite setIC divgI -card_quotient // -quotientMidl NG_AG.
+ rewrite card_quotient -?divgS //= norm_joinEl //.
+ by rewrite coprime_cardMg 1?coprime_sym // mulnK.
+have hallGN: Hall N (G :&: N).
+ by rewrite /Hall -divgS subsetIr //= iGN_A (coprimeSg _ coGA) ?subsetIl.
+case/splitsP: {hallGN nGN_N}(SchurZassenhaus_split hallGN nGN_N) => B.
+case/complP=> trBGN defN.
+have{trBGN iGN_A} oBA: #|B| = #|A|.
+ by rewrite -iGN_A -{1}defN (TI_cardMg trBGN) mulKn.
+have sBN: B \subset N by rewrite -defN mulG_subr.
+case: (SchurZassenhaus_trans_sol solG nGA _ coGA oBA) => [|x Gx defB].
+ by rewrite -(normC nGA) -norm_joinEl // -NG_AG -(mul1g B) mulgSS ?sub1G.
+exists (H :^ x^-1)%G; first by rewrite pHallJ ?groupV.
+apply/subsetP=> y Ay; have: y ^ x \in B by rewrite defB memJ_conjg.
+move/(subsetP sBN)=> /setIP[_ /normP nHyx].
+by apply/normP; rewrite -conjsgM conjgCV invgK conjsgM nHyx.
+Qed.
+
+(* This is B & G, Proposition 1.5(c) *)
+Proposition coprime_Hall_trans A G H1 H2 :
+ A \subset 'N(G) -> coprime #|G| #|A| -> solvable G ->
+ pi.-Hall(G) H1 -> A \subset 'N(H1) ->
+ pi.-Hall(G) H2 -> A \subset 'N(H2) ->
+ exists2 x, x \in 'C_G(A) & H1 :=: H2 :^ x.
+Proof.
+move: H1 => H nGA coGA solG hallH nHA hallH2.
+have{H2 hallH2} [x Gx -> nH1xA] := Hall_trans solG hallH2 hallH.
+have sG_AG: G \subset A <*> G by rewrite -{1}genGid genS ?subsetUr.
+have nG_AG: A <*> G \subset 'N(G) by rewrite gen_subG subUset nGA normG.
+pose N := 'N_(A <*> G)(H)%G.
+have nGN: N \subset 'N(G) by rewrite subIset ?nG_AG.
+have nGN_N: G :&: N <| N.
+ apply/normalP; rewrite subsetIr; split=> // y Ny.
+ by rewrite conjIg (normP _) // (subsetP nGN, conjGid).
+have NG_AG : G * N = A <*> G.
+ by apply: Hall_Frattini_arg hallH => //; exact/andP.
+have iGN_A: #|N : G :&: N| = #|A|.
+ rewrite -card_quotient //; last by case/andP: nGN_N.
+ rewrite (card_isog (second_isog nGN)) /= -quotientMidr (normC nGN) NG_AG.
+ rewrite card_quotient // -divgS //= joingC norm_joinEr //.
+ by rewrite coprime_cardMg // mulnC mulnK.
+have solGN: solvable (G :&: N) by apply: solvableS solG; exact: subsetIl.
+have oAxA: #|A :^ x^-1| = #|A| by exact: cardJg.
+have sAN: A \subset N by rewrite subsetI -{1}genGid genS // subsetUl.
+have nGNA: A \subset 'N(G :&: N).
+ by apply/normsP=> y ?; rewrite conjIg (normsP nGA) ?(conjGid, subsetP sAN).
+have coGNA: coprime #|G :&: N| #|A| := coprimeSg (subsetIl _ _) coGA.
+case: (SchurZassenhaus_trans_sol solGN nGNA _ coGNA oAxA) => [|y GNy defAx].
+ have ->: (G :&: N) * A = N.
+ apply/eqP; rewrite eqEcard -{2}(mulGid N) mulgSS ?subsetIr //=.
+ by rewrite coprime_cardMg // -iGN_A Lagrange ?subsetIr.
+ rewrite sub_conjgV conjIg -normJ subsetI conjGid ?joing_subl //.
+ by rewrite mem_gen // inE Gx orbT.
+case/setIP: GNy => Gy; case/setIP=> _; move/normP=> nHy.
+exists (y * x)^-1.
+ rewrite -coprime_norm_cent // groupV inE groupM //=; apply/normP.
+ by rewrite conjsgM -defAx conjsgKV.
+by apply: val_inj; rewrite /= -{2}nHy -(conjsgM _ y) conjsgK.
+Qed.
+
+(* A complement to the above: 'C(A) acts on 'Nby(A) *)
+Lemma norm_conj_cent A G x : x \in 'C(A) ->
+ (A \subset 'N(G :^ x)) = (A \subset 'N(G)).
+Proof. by move=> cAx; rewrite norm_conj_norm ?(subsetP (cent_sub A)). Qed.
+
+(* Strongest version of the centraliser lemma -- not found in textbooks! *)
+(* Obviously, the solvability condition could be removed once we have the *)
+(* Odd Order Theorem. *)
+Lemma strongest_coprime_quotient_cent A G H :
+ let R := H :&: [~: G, A] in
+ A \subset 'N(H) -> R \subset G -> coprime #|R| #|A| ->
+ solvable R || solvable A ->
+ 'C_G(A) / H = 'C_(G / H)(A / H).
+Proof.
+move=> R nHA sRG coRA solRA.
+have nRA: A \subset 'N(R) by rewrite normsI ?commg_normr.
+apply/eqP; rewrite eqEsubset subsetI morphimS ?subsetIl //=.
+rewrite (subset_trans _ (morphim_cent _ _)) ?morphimS ?subsetIr //=.
+apply/subsetP=> _ /setIP[/morphimP[x Nx Gx ->] cAHx].
+have{cAHx} cAxR y: y \in A -> [~ x, y] \in R.
+ move=> Ay; have Ny: y \in 'N(H) by exact: subsetP Ay.
+ rewrite inE mem_commg // andbT coset_idr ?groupR // morphR //=.
+ by apply/eqP; apply/commgP; apply: (centP cAHx); rewrite mem_quotient.
+have AxRA: A :^ x \subset R * A.
+ apply/subsetP=> _ /imsetP[y Ay ->].
+ rewrite -normC // -(mulKVg y (y ^ x)) -commgEl mem_mulg //.
+ by rewrite -groupV invg_comm cAxR.
+have [y Ry def_Ax]: exists2 y, y \in R & A :^ x = A :^ y.
+ have oAx: #|A :^ x| = #|A| by rewrite cardJg.
+ have [solR | solA] := orP solRA; first exact: SchurZassenhaus_trans_sol.
+ by apply: SchurZassenhaus_trans_actsol; rewrite // joingC norm_joinEr.
+rewrite -imset_coset; apply/imsetP; exists (x * y^-1); last first.
+ by rewrite conjgCV mkerl // ker_coset memJ_norm groupV; case/setIP: Ry.
+rewrite /= inE groupMl // ?(groupV, subsetP sRG) //=.
+apply/centP=> z Az; apply/commgP/eqP/set1P.
+rewrite -[[set 1]](coprime_TIg coRA) inE {1}commgEl commgEr /= -/R.
+rewrite invMg -mulgA invgK groupMl // conjMg mulgA -commgEl.
+rewrite groupMl ?cAxR // memJ_norm ?(groupV, subsetP nRA) // Ry /=.
+by rewrite groupMr // conjVg groupV conjgM -mem_conjg -def_Ax memJ_conjg.
+Qed.
+
+(* A weaker but more practical version, still stronger than the usual form *)
+(* (viz. Aschbacher 18.7.4), similar to the one needed in Aschbacher's *)
+(* proof of Thompson factorization. Note that the coprime and solvability *)
+(* assumptions could be further weakened to H :&: G (and hence become *)
+(* trivial if H and G are TI). However, the assumption that A act on G is *)
+(* needed in this case. *)
+Lemma coprime_norm_quotient_cent A G H :
+ A \subset 'N(G) -> A \subset 'N(H) -> coprime #|H| #|A| -> solvable H ->
+ 'C_G(A) / H = 'C_(G / H)(A / H).
+Proof.
+move=> nGA nHA coHA solH; have sRH := subsetIl H [~: G, A].
+rewrite strongest_coprime_quotient_cent ?(coprimeSg sRH) 1?(solvableS sRH) //.
+by rewrite subIset // commg_subl nGA orbT.
+Qed.
+
+(* A useful consequence (similar to Ex. 6.1 in Aschbacher) of the stronger *)
+(* theorem. *)
+Lemma coprime_cent_mulG A G H :
+ A \subset 'N(G) -> A \subset 'N(H) -> G \subset 'N(H) ->
+ coprime #|H| #|A| -> solvable H ->
+ 'C_(H * G)(A) = 'C_H(A) * 'C_G(A).
+Proof.
+move=> nHA nGA nHG coHA solH; rewrite -norm_joinEr //.
+have nsHG: H <| H <*> G by rewrite /normal joing_subl join_subG normG.
+rewrite -{2}(setIidPr (normal_sub nsHG)) setIAC.
+rewrite group_modr ?setSI ?joing_subr //=; symmetry; apply/setIidPl.
+rewrite -quotientSK ?subIset 1?normal_norm //.
+by rewrite !coprime_norm_quotient_cent ?normsY //= norm_joinEr ?quotientMidl.
+Qed.
+
+(* Another special case of the strong coprime quotient lemma; not found in *)
+(* textbooks, but nevertheless used implicitly throughout B & G, sometimes *)
+(* justified by switching to external action. *)
+Lemma quotient_TI_subcent K G H :
+ G \subset 'N(K) -> G \subset 'N(H) -> K :&: H = 1 ->
+ 'C_K(G) / H = 'C_(K / H)(G / H).
+Proof.
+move=> nGK nGH tiKH.
+have tiHR: H :&: [~: K, G] = 1.
+ by apply/trivgP; rewrite /= setIC -tiKH setSI ?commg_subl.
+apply: strongest_coprime_quotient_cent; rewrite ?tiHR ?sub1G ?solvable1 //.
+by rewrite cards1 coprime1n.
+Qed.
+
+(* This is B & G, Proposition 1.5(d): the more traditional form of the lemma *)
+(* above, with the assumption H <| G weakened to H \subset G. The stronger *)
+(* coprime and solvability assumptions are easier to satisfy in practice. *)
+Proposition coprime_quotient_cent A G H :
+ H \subset G -> A \subset 'N(H) -> coprime #|G| #|A| -> solvable G ->
+ 'C_G(A) / H = 'C_(G / H)(A / H).
+Proof.
+move=> sHG nHA coGA solG.
+have sRG: H :&: [~: G, A] \subset G by rewrite subIset ?sHG.
+by rewrite strongest_coprime_quotient_cent ?(coprimeSg sRG) 1?(solvableS sRG).
+Qed.
+
+(* This is B & G, Proposition 1.5(e). *)
+Proposition coprime_comm_pcore A G K :
+ A \subset 'N(G) -> coprime #|G| #|A| -> solvable G ->
+ pi^'.-Hall(G) K -> K \subset 'C_G(A) ->
+ [~: G, A] \subset 'O_pi(G).
+Proof.
+move=> nGA coGA solG hallK cKA.
+case: (coprime_Hall_exists nGA) => // H hallH nHA.
+have sHG: H \subset G by case/andP: hallH.
+have sKG: K \subset G by case/andP: hallK.
+have coKH: coprime #|K| #|H|.
+ case/and3P: hallH=> _ piH _; case/and3P: hallK => _ pi'K _.
+ by rewrite coprime_sym (pnat_coprime piH pi'K).
+have defG: G :=: K * H.
+ apply/eqP; rewrite eq_sym eqEcard coprime_cardMg //.
+ rewrite -{1}(mulGid G) mulgSS //= (card_Hall hallH) (card_Hall hallK).
+ by rewrite mulnC partnC.
+have sGA_H: [~: G, A] \subset H.
+ rewrite gen_subG defG; apply/subsetP=> xya; case/imset2P=> xy a.
+ case/imset2P=> x y Kx Hy -> Aa -> {xya xy}.
+ rewrite commMgJ (([~ x, a] =P 1) _) ?(conj1g, mul1g).
+ by rewrite groupMl ?groupV // memJ_norm ?(subsetP nHA).
+ rewrite subsetI sKG in cKA; apply/commgP; exact: (centsP cKA).
+apply: pcore_max; last first.
+ by rewrite /(_ <| G) /= commg_norml commGC commg_subr nGA.
+by case/and3P: hallH => _ piH _; apply: pgroupS piH.
+Qed.
+
+End InternalAction.
+
+(* This is B & G, Proposition 1.5(b). *)
+Proposition coprime_Hall_subset pi (gT : finGroupType) (A G X : {group gT}) :
+ A \subset 'N(G) -> coprime #|G| #|A| -> solvable G ->
+ X \subset G -> pi.-group X -> A \subset 'N(X) ->
+ exists H : {group gT}, [/\ pi.-Hall(G) H, A \subset 'N(H) & X \subset H].
+Proof.
+move: {2}_.+1 (ltnSn #|G|) => n.
+elim: n => // n IHn in gT A G X * => leGn nGA coGA solG sXG piX nXA.
+have [G1 | ntG] := eqsVneq G 1.
+ case: (coprime_Hall_exists pi nGA) => // H hallH nHA.
+ by exists H; split; rewrite // (subset_trans sXG) // G1 sub1G.
+have sG_AG: G \subset A <*> G by rewrite joing_subr.
+have sA_AG: A \subset A <*> G by rewrite joing_subl.
+have nG_AG: A <*> G \subset 'N(G) by rewrite join_subG nGA normG.
+have nsG_AG: G <| A <*> G by exact/andP.
+case: (solvable_norm_abelem solG nsG_AG) => // M [sMG nsMAG ntM].
+have{nsMAG} [nMA nMG]: A \subset 'N(M) /\ G \subset 'N(M).
+ by apply/andP; rewrite -join_subG normal_norm.
+have nMX: X \subset 'N(M) by exact: subset_trans nMG.
+case/is_abelemP=> p pr_p; case/and3P=> pM cMM _.
+have: #|G / M| < n by rewrite (leq_trans (ltn_quotient _ _)).
+move/(IHn _ (A / M)%G _ (X / M)%G); rewrite !(quotient_norms, quotientS) //.
+rewrite !(coprime_morph, quotient_sol, morphim_pgroup) //.
+case=> //= Hq []; case/and3P=> sHGq piHq pi'Hq' nHAq sXHq.
+case/inv_quotientS: (sHGq) => [|HM defHM sMHM sHMG]; first exact/andP.
+have nMHM := subset_trans sHMG nMG.
+have{sXHq} sXHM: X \subset HM by rewrite -(quotientSGK nMX) -?defHM.
+have{pi'Hq' sHGq} pi'HM': pi^'.-nat #|G : HM|.
+ move: pi'Hq'; rewrite -!divgS // defHM !card_quotient //.
+ by rewrite -(divnMl (cardG_gt0 M)) !Lagrange.
+have{nHAq} nHMA: A \subset 'N(HM).
+ by rewrite -(quotientSGK nMA) ?normsG ?quotient_normG -?defHM //; exact/andP.
+case/orP: (orbN (p \in pi)) => pi_p.
+ exists HM; split=> //; apply/and3P; split; rewrite /pgroup //.
+ by rewrite -(Lagrange sMHM) pnat_mul -card_quotient // -defHM (pi_pnat pM).
+case: (ltnP #|HM| #|G|) => [ltHG | leGHM {n IHn leGn}].
+ case: (IHn _ A HM X (leq_trans ltHG leGn)) => // [||H [hallH nHA sXH]].
+ - exact: coprimeSg coGA.
+ - exact: solvableS solG.
+ case/and3P: hallH => sHHM piH pi'H'.
+ have sHG: H \subset G by exact: subset_trans sHMG.
+ exists H; split=> //; apply/and3P; split=> //.
+ rewrite -divgS // -(Lagrange sHMG) -(Lagrange sHHM) -mulnA mulKn //.
+ by rewrite pnat_mul pi'H'.
+have{leGHM nHMA sHMG sMHM sXHM pi'HM'} eqHMG: HM = G.
+ by apply/eqP; rewrite -val_eqE eqEcard sHMG.
+have pi'M: pi^'.-group M by rewrite /pgroup (pi_pnat pM).
+have{HM Hq nMHM defHM eqHMG piHq} hallM: pi^'.-Hall(G) M.
+ apply/and3P; split; rewrite // /pgroup pnatNK.
+ by rewrite -card_quotient // -eqHMG -defHM.
+case: (coprime_Hall_exists pi nGA) => // H hallH nHA.
+pose XM := (X <*> M)%G; pose Y := (H :&: XM)%G.
+case/and3P: (hallH) => sHG piH _.
+have sXXM: X \subset XM by rewrite joing_subl.
+have co_pi_M (B : {group gT}): pi.-group B -> coprime #|B| #|M|.
+ by move=> piB; rewrite (pnat_coprime piB).
+have hallX: pi.-Hall(XM) X.
+ rewrite /pHall piX sXXM -divgS //= norm_joinEl //.
+ by rewrite coprime_cardMg ?co_pi_M // mulKn.
+have sXMG: XM \subset G by rewrite join_subG sXG.
+have hallY: pi.-Hall(XM) Y.
+ have sYXM: Y \subset XM by rewrite subsetIr.
+ have piY: pi.-group Y by apply: pgroupS piH; exact: subsetIl.
+ rewrite /pHall sYXM piY -divgS // -(_ : Y * M = XM).
+ by rewrite coprime_cardMg ?co_pi_M // mulKn //.
+ rewrite /= setIC group_modr ?joing_subr //=; apply/setIidPl.
+ rewrite ((H * M =P G) _) // eqEcard mul_subG //= coprime_cardMg ?co_pi_M //.
+ by rewrite (card_Hall hallM) (card_Hall hallH) partnC.
+have nXMA: A \subset 'N(XM) by rewrite normsY.
+have:= coprime_Hall_trans nXMA _ _ hallX nXA hallY.
+rewrite !(coprimeSg sXMG, solvableS sXMG, normsI) //.
+case=> // x /setIP[XMx cAx] ->.
+exists (H :^ x)%G; split; first by rewrite pHallJ ?(subsetP sXMG).
+ by rewrite norm_conj_cent.
+by rewrite conjSg subsetIl.
+Qed.
+
+Section ExternalAction.
+
+Variables (pi : nat_pred) (aT gT : finGroupType).
+Variables (A : {group aT}) (G : {group gT}) (to : groupAction A G).
+
+Section FullExtension.
+
+Local Notation inA := (sdpair2 to).
+Local Notation inG := (sdpair1 to).
+Local Notation A' := (inA @* gval A).
+Local Notation G' := (inG @* gval G).
+Let injG : 'injm inG := injm_sdpair1 _.
+Let injA : 'injm inA := injm_sdpair2 _.
+
+Hypotheses (coGA : coprime #|G| #|A|) (solG : solvable G).
+
+Lemma external_action_im_coprime : coprime #|G'| #|A'|.
+Proof. by rewrite !card_injm. Qed.
+
+Let coGA' := external_action_im_coprime.
+
+Let solG' : solvable G' := morphim_sol _ solG.
+
+Let nGA' := im_sdpair_norm to.
+
+Lemma ext_coprime_Hall_exists :
+ exists2 H : {group gT}, pi.-Hall(G) H & [acts A, on H | to].
+Proof.
+have [H' hallH' nHA'] := coprime_Hall_exists pi nGA' coGA' solG'.
+have sHG' := pHall_sub hallH'.
+exists (inG @*^-1 H')%G => /=.
+ by rewrite -(morphim_invmE injG) -{1}(im_invm injG) morphim_pHall.
+by rewrite actsEsd ?morphpreK // subsetIl.
+Qed.
+
+Lemma ext_coprime_Hall_trans (H1 H2 : {group gT}) :
+ pi.-Hall(G) H1 -> [acts A, on H1 | to] ->
+ pi.-Hall(G) H2 -> [acts A, on H2 | to] ->
+ exists2 x, x \in 'C_(G | to)(A) & H1 :=: H2 :^ x.
+Proof.
+move=> hallH1 nH1A hallH2 nH2A.
+have sH1G := pHall_sub hallH1; have sH2G := pHall_sub hallH2.
+rewrite !actsEsd // in nH1A nH2A.
+have hallH1': pi.-Hall(G') (inG @* H1) by rewrite morphim_pHall.
+have hallH2': pi.-Hall(G') (inG @* H2) by rewrite morphim_pHall.
+have [x'] := coprime_Hall_trans nGA' coGA' solG' hallH1' nH1A hallH2' nH2A.
+case/setIP=> /= Gx' cAx' /eqP defH1; pose x := invm injG x'.
+have Gx: x \in G by rewrite -(im_invm injG) mem_morphim.
+have def_x': x' = inG x by rewrite invmK.
+exists x; first by rewrite inE Gx gacentEsd mem_morphpre /= -?def_x'.
+apply/eqP; move: defH1; rewrite def_x' /= -morphimJ //=.
+by rewrite !eqEsubset !injmSK // conj_subG.
+Qed.
+
+Lemma ext_norm_conj_cent (H : {group gT}) x :
+ H \subset G -> x \in 'C_(G | to)(A) ->
+ [acts A, on H :^ x | to] = [acts A, on H | to].
+Proof.
+move=> sHG /setIP[Gx].
+rewrite gacentEsd !actsEsd ?conj_subG ?morphimJ // 2!inE Gx /=.
+exact: norm_conj_cent.
+Qed.
+
+Lemma ext_coprime_Hall_subset (X : {group gT}) :
+ X \subset G -> pi.-group X -> [acts A, on X | to] ->
+ exists H : {group gT},
+ [/\ pi.-Hall(G) H, [acts A, on H | to] & X \subset H].
+Proof.
+move=> sXG piX; rewrite actsEsd // => nXA'.
+case: (coprime_Hall_subset nGA' coGA' solG' _ (morphim_pgroup _ piX) nXA').
+ exact: morphimS.
+move=> H' /= [piH' nHA' sXH']; have sHG' := pHall_sub piH'.
+exists (inG @*^-1 H')%G; rewrite actsEsd ?subsetIl ?morphpreK // nHA'.
+rewrite -sub_morphim_pre //= sXH'; split=> //.
+by rewrite -(morphim_invmE injG) -{1}(im_invm injG) morphim_pHall.
+Qed.
+
+End FullExtension.
+
+(* We only prove a weaker form of the coprime group action centraliser *)
+(* lemma, because it is more convenient in practice to make G the range *)
+(* of the action, whence G both contains H and is stable under A. *)
+(* However we do restrict the coprime/solvable assumptions to H, and *)
+(* we do not require that G normalize H. *)
+Lemma ext_coprime_quotient_cent (H : {group gT}) :
+ H \subset G -> [acts A, on H | to] -> coprime #|H| #|A| -> solvable H ->
+ 'C_(|to)(A) / H = 'C_(|to / H)(A).
+Proof.
+move=> sHG nHA coHA solH; pose N := 'N_G(H).
+have nsHN: H <| N by rewrite normal_subnorm.
+have [sHN nHn] := andP nsHN.
+have sNG: N \subset G by exact: subsetIl.
+have nNA: {acts A, on group N | to}.
+ split; rewrite // actsEsd // injm_subnorm ?injm_sdpair1 //=.
+ by rewrite normsI ?norms_norm ?im_sdpair_norm -?actsEsd.
+rewrite -!(gacentIdom _ A) -quotientInorm -gacentIim setIAC.
+rewrite -(gacent_actby nNA) gacentEsd -morphpreIim /= -/N.
+have:= (injm_sdpair1 <[nNA]>, injm_sdpair2 <[nNA]>).
+set inG := sdpair1 _; set inA := sdpair2 _ => [[injG injA]].
+set G' := inG @* N; set A' := inA @* A; pose H' := inG @* H.
+have defN: 'N(H | to) = A by apply/eqP; rewrite eqEsubset subsetIl.
+have def_Dq: qact_dom to H = A by rewrite qact_domE.
+have sAq: A \subset qact_dom to H by rewrite def_Dq.
+rewrite {2}def_Dq -(gacent_ract _ sAq); set to_q := (_ \ _)%gact.
+have:= And3 (sdprod_sdpair to_q) (injm_sdpair1 to_q) (injm_sdpair2 to_q).
+rewrite gacentEsd; set inAq := sdpair2 _; set inGq := sdpair1 _ => /=.
+set Gq := inGq @* _; set Aq := inAq @* _ => [[q_d iGq iAq]].
+have nH': 'N(H') = setT.
+ apply/eqP; rewrite -subTset -im_sdpair mulG_subG morphim_norms //=.
+ by rewrite -actsEsd // acts_actby subxx /= (setIidPr sHN).
+have: 'dom (coset H' \o inA \o invm iAq) = Aq.
+ by rewrite ['dom _]morphpre_invm /= nH' morphpreT.
+case/domP=> qA [def_qA ker_qA _ im_qA].
+have{coHA} coHA': coprime #|H'| #|A'| by rewrite !card_injm.
+have{ker_qA} injAq: 'injm qA.
+ rewrite {}ker_qA !ker_comp ker_coset morphpre_invm -morphpreIim /= setIC.
+ by rewrite coprime_TIg // -kerE (trivgP injA) morphim1.
+have{im_qA} im_Aq : qA @* Aq = A' / H'.
+ by rewrite {}im_qA !morphim_comp im_invm.
+have: 'dom (quotm (sdpair1_morphism <[nNA]>) nsHN \o invm iGq) = Gq.
+ by rewrite ['dom _]morphpre_invm /= quotientInorm.
+case/domP=> qG [def_qG ker_qG _ im_qG].
+have{ker_qG} injGq: 'injm qG.
+ rewrite {}ker_qG ker_comp ker_quotm morphpre_invm (trivgP injG).
+ by rewrite quotient1 morphim1.
+have im_Gq: qG @* Gq = G' / H'.
+ rewrite {}im_qG morphim_comp im_invm morphim_quotm //= -/inG -/H'.
+ by rewrite -morphimIdom setIAC setIid.
+have{def_qA def_qG} q_J : {in Gq & Aq, morph_act 'J 'J qG qA}.
+ move=> x' a'; case/morphimP=> Hx; case/morphimP=> x nHx Gx -> GHx ->{Hx x'}.
+ case/morphimP=> a _ Aa ->{a'} /=; rewrite -/inAq -/inGq.
+ rewrite !{}def_qG {}def_qA /= !invmE // -sdpair_act //= -/inG -/inA.
+ have Nx: x \in N by rewrite inE Gx.
+ have Nxa: to x a \in N by case: (nNA); move/acts_act->.
+ have [Gxa nHxa] := setIP Nxa.
+ rewrite invmE qactE ?quotmE ?mem_morphim ?def_Dq //=.
+ by rewrite -morphJ /= ?nH' ?inE // -sdpair_act //= actbyE.
+pose q := sdprodm q_d q_J.
+have{injAq injGq} injq: 'injm q.
+ rewrite injm_sdprodm injAq injGq /= {}im_Aq {}im_Gq -/Aq .
+ by rewrite -quotientGI ?im_sdpair_TI ?morphimS //= quotient1.
+rewrite -[inGq @*^-1 _]morphpreIim -/Gq.
+have sC'G: inG @*^-1 'C_G'(A') \subset G by rewrite !subIset ?subxx.
+rewrite -[_ / _](injmK iGq) ?quotientS //= -/inGq; congr (_ @*^-1 _).
+apply: (injm_morphim_inj injq); rewrite 1?injm_subcent ?subsetT //= -/q.
+rewrite 2?morphim_sdprodml ?morphimS //= im_Gq.
+rewrite morphim_sdprodmr ?morphimS //= im_Aq.
+rewrite {}im_qG morphim_comp morphim_invm ?morphimS //.
+rewrite morphim_quotm morphpreK ?subsetIl //= -/H'.
+rewrite coprime_norm_quotient_cent ?im_sdpair_norm ?nH' ?subsetT //=.
+exact: morphim_sol.
+Qed.
+
+End ExternalAction.
+
+Section SylowSolvableAct.
+
+Variables (gT : finGroupType) (p : nat).
+Implicit Types A B G X : {group gT}.
+
+Lemma sol_coprime_Sylow_exists A G :
+ solvable A -> A \subset 'N(G) -> coprime #|G| #|A| ->
+ exists2 P : {group gT}, p.-Sylow(G) P & A \subset 'N(P).
+Proof.
+move=> solA nGA coGA; pose AG := A <*> G.
+have nsG_AG: G <| AG by rewrite /normal joing_subr join_subG nGA normG.
+have [sG_AG nG_AG]:= andP nsG_AG.
+have [P sylP] := Sylow_exists p G; pose N := 'N_AG(P); pose NG := G :&: N.
+have nGN: N \subset 'N(G) by rewrite subIset ?nG_AG.
+have sNG_G: NG \subset G := subsetIl G N.
+have nsNG_N: NG <| N by rewrite /normal subsetIr normsI ?normG.
+have defAG: G * N = AG := Frattini_arg nsG_AG sylP.
+have oA : #|A| = #|N| %/ #|NG|.
+ rewrite /NG setIC divgI -card_quotient // -quotientMidl defAG.
+ rewrite card_quotient -?divgS //= norm_joinEl //.
+ by rewrite coprime_cardMg 1?coprime_sym // mulnK.
+have: [splits N, over NG].
+ rewrite SchurZassenhaus_split // /Hall -divgS subsetIr //.
+ by rewrite -oA (coprimeSg sNG_G).
+case/splitsP=> B; case/complP=> tNG_B defN.
+have [nPB]: B \subset 'N(P) /\ B \subset AG.
+ by apply/andP; rewrite andbC -subsetI -/N -defN mulG_subr.
+case/SchurZassenhaus_trans_actsol => // [|x Gx defB].
+ by rewrite oA -defN TI_cardMg // mulKn.
+exists (P :^ x^-1)%G; first by rewrite pHallJ ?groupV.
+by rewrite normJ -sub_conjg -defB.
+Qed.
+
+Lemma sol_coprime_Sylow_trans A G :
+ solvable A -> A \subset 'N(G) -> coprime #|G| #|A| ->
+ [transitive 'C_G(A), on [set P in 'Syl_p(G) | A \subset 'N(P)] | 'JG].
+Proof.
+move=> solA nGA coGA; pose AG := A <*> G; set FpA := finset _.
+have nG_AG: AG \subset 'N(G) by rewrite join_subG nGA normG.
+have [P sylP nPA] := sol_coprime_Sylow_exists solA nGA coGA.
+pose N := 'N_AG(P); have sAN: A \subset N by rewrite subsetI joing_subl.
+have trNPA: A :^: AG ::&: N = A :^: N.
+ pose NG := 'N_G(P); have sNG_G : NG \subset G := subsetIl _ _.
+ have nNGA: A \subset 'N(NG) by rewrite normsI ?norms_norm.
+ apply/setP=> Ax; apply/setIdP/imsetP=> [[]|[x Nx ->{Ax}]]; last first.
+ by rewrite conj_subG //; case/setIP: Nx => AGx; rewrite mem_imset.
+ have ->: N = A <*> NG by rewrite /N /AG !norm_joinEl // -group_modl.
+ have coNG_A := coprimeSg sNG_G coGA; case/imsetP=> x AGx ->{Ax}.
+ case/SchurZassenhaus_trans_actsol; rewrite ?cardJg // => y Ny /= ->.
+ by exists y; rewrite // mem_gen 1?inE ?Ny ?orbT.
+have{trNPA}: [transitive 'N_AG(A), on FpA | 'JG].
+ have ->: FpA = 'Fix_('Syl_p(G) | 'JG)(A).
+ by apply/setP=> Q; rewrite 4!inE afixJG.
+ have SylP : P \in 'Syl_p(G) by rewrite inE.
+ apply/(trans_subnorm_fixP _ SylP); rewrite ?astab1JG //.
+ rewrite (atrans_supgroup _ (Syl_trans _ _)) ?joing_subr //= -/AG.
+ by apply/actsP=> x /= AGx Q /=; rewrite !inE -{1}(normsP nG_AG x) ?pHallJ2.
+rewrite {1}/AG norm_joinEl // -group_modl ?normG ?coprime_norm_cent //=.
+rewrite -cent_joinEr ?subsetIr // => trC_FpA.
+have FpA_P: P \in FpA by rewrite !inE sylP.
+apply/(subgroup_transitiveP FpA_P _ trC_FpA); rewrite ?joing_subr //=.
+rewrite astab1JG cent_joinEr ?subsetIr // -group_modl // -mulgA.
+by congr (_ * _); rewrite mulSGid ?subsetIl.
+Qed.
+
+Lemma sol_coprime_Sylow_subset A G X :
+ A \subset 'N(G) -> coprime #|G| #|A| -> solvable A ->
+ X \subset G -> p.-group X -> A \subset 'N(X) ->
+ exists P : {group gT}, [/\ p.-Sylow(G) P, A \subset 'N(P) & X \subset P].
+Proof.
+move=> nGA coGA solA sXG pX nXA.
+pose nAp (Q : {group gT}) := [&& p.-group Q, Q \subset G & A \subset 'N(Q)].
+have: nAp X by exact/and3P.
+case/maxgroup_exists=> R; case/maxgroupP; case/and3P=> pR sRG nRA maxR sXR.
+have [P sylP sRP]:= Sylow_superset sRG pR.
+suffices defP: P :=: R by exists P; rewrite sylP defP.
+case/and3P: sylP => sPG pP _; apply: (nilpotent_sub_norm (pgroup_nil pP)) => //.
+pose N := 'N_G(R); have{sPG} sPN_N: 'N_P(R) \subset N by exact: setSI.
+apply: norm_sub_max_pgroup (pgroupS (subsetIl _ _) pP) sPN_N (subsetIr _ _).
+have nNA: A \subset 'N(N) by rewrite normsI ?norms_norm.
+have coNA: coprime #|N| #|A| by apply: coprimeSg coGA; rewrite subsetIl.
+have{solA coNA} [Q sylQ nQA] := sol_coprime_Sylow_exists solA nNA coNA.
+suffices defQ: Q :=: R by rewrite max_pgroup_Sylow -{2}defQ.
+apply: maxR; first by apply/and3P; case/and3P: sylQ; rewrite subsetI; case/andP.
+by apply: normal_sub_max_pgroup (Hall_max sylQ) pR _; rewrite normal_subnorm.
+Qed.
+
+End SylowSolvableAct.
diff --git a/mathcomp/solvable/jordanholder.v b/mathcomp/solvable/jordanholder.v
new file mode 100644
index 0000000..bc7bec2
--- /dev/null
+++ b/mathcomp/solvable/jordanholder.v
@@ -0,0 +1,681 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path choice fintype.
+Require Import bigop finset fingroup morphism automorphism quotient action.
+Require Import gseries.
+
+(******************************************************************************)
+(* This files establishes Jordan-Holder theorems for finite groups. These *)
+(* theorems state the uniqueness up to permutation and isomorphism for the *)
+(* series of quotient built from the successive elements of any composition *)
+(* series of the same group. These quotients are also called factors of the *)
+(* composition series. To avoid the heavy use of highly polymorphic lists *)
+(* describing these quotient series, we introduce sections. *)
+(* This library defines: *)
+(* (G1 / G2)%sec == alias for the pair (G1, G2) of groups in the same *)
+(* finGroupType, coerced to the actual quotient group*)
+(* group G1 / G2. We call this pseudo-quotient a *)
+(* section of G1 and G2. *)
+(* section_isog s1 s2 == s1 and s2 respectively coerce to isomorphic *)
+(* quotient groups. *)
+(* section_repr s == canonical representative of the isomorphism class *)
+(* of the section s. *)
+(* mksrepr G1 G2 == canonical representative of the isomorphism class *)
+(* of (G1 / G2)%sec. *)
+(* mkfactors G s == if s is [:: s1, s2, ..., sn], constructs the list *)
+(* [:: mksrepr G s1, mksrepr s1 s2, ..., mksrepr sn-1 sn] *)
+(* comps G s == s is a composition series for G i.e. s is a *)
+(* decreasing sequence of subgroups of G *)
+(* in which two adjacent elements are maxnormal one *)
+(* in the other and the last element of s is 1. *)
+(* Given aT and rT two finGroupTypes, (D : {group rT}), (A : {group aT}) and *)
+(* (to : groupAction A D) an external action. *)
+(* maxainv to B C == C is a maximal proper normal subgroup of B *)
+(* invariant by (the external action of A via) to. *)
+(* asimple to B == the maximal proper normal subgroup of B invariant *)
+(* by the external action to is trivial. *)
+(* acomps to G s == s is a composition series for G invariant by to, *)
+(* i.e. s is a decreasing sequence of subgroups of G *)
+(* in which two adjacent elements are maximally *)
+(* invariant by to one in the other and the *)
+(* last element of s is 1. *)
+(* We prove two versions of the result: *)
+(* - JordanHolderUniqueness establishes the uniqueness up to permutation *)
+(* and isomorphism of the lists of factors in composition series of a *)
+(* given group. *)
+(* - StrongJordanHolderUniqueness extends the result to composition series *)
+(* invariant by an external group action. *)
+(* See also "The Rooster and the Butterflies", proceedings of Calculemus 2013,*)
+(* by Assia Mahboubi. *)
+(******************************************************************************)
+
+
+Import GroupScope.
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Inductive section (gT : finGroupType) := GSection of {group gT} * {group gT}.
+
+Delimit Scope section_scope with sec.
+Bind Scope section_scope with section.
+
+Definition mkSec (gT : finGroupType) (G1 G2 : {group gT}) := GSection (G1, G2).
+
+Infix "/" := mkSec : section_scope.
+
+Coercion pair_of_section gT (s : section gT) := let: GSection u := s in u.
+
+Coercion quotient_of_section gT (s : section gT) : GroupSet.sort _ := s.1 / s.2.
+
+Coercion section_group gT (s : section gT) : {group (coset_of s.2)} :=
+ Eval hnf in [group of s].
+
+Section Sections.
+
+Variables (gT : finGroupType).
+Implicit Types (G : {group gT}) (s : section gT).
+
+Canonical section_subType := Eval hnf in [newType for @pair_of_section gT].
+Definition section_eqMixin := Eval hnf in [eqMixin of section gT by <:].
+Canonical section_eqType := Eval hnf in EqType (section gT) section_eqMixin.
+Definition section_choiceMixin := [choiceMixin of section gT by <:].
+Canonical section_choiceType :=
+ Eval hnf in ChoiceType (section gT) section_choiceMixin.
+Definition section_countMixin := [countMixin of section gT by <:].
+Canonical section_countType :=
+ Eval hnf in CountType (section gT) section_countMixin.
+Canonical section_subCountType := Eval hnf in [subCountType of section gT].
+Definition section_finMixin := [finMixin of section gT by <:].
+Canonical section_finType := Eval hnf in FinType (section gT) section_finMixin.
+Canonical section_subFinType := Eval hnf in [subFinType of section gT].
+Canonical section_group.
+
+(* Isomorphic sections *)
+
+Definition section_isog := [rel x y : section gT | x \isog y].
+
+(* A witness of the isomorphism class of a section *)
+Definition section_repr s := odflt (1 / 1)%sec (pick (section_isog ^~ s)).
+
+Definition mksrepr G1 G2 := section_repr (mkSec G1 G2).
+
+Lemma section_reprP s : section_repr s \isog s.
+Proof.
+by rewrite /section_repr; case: pickP => //= /(_ s); rewrite isog_refl.
+Qed.
+
+Lemma section_repr_isog s1 s2 :
+ s1 \isog s2 -> section_repr s1 = section_repr s2.
+Proof.
+by move=> iso12; congr (odflt _ _); apply: eq_pick => s; exact: isog_transr.
+Qed.
+
+Definition mkfactors (G : {group gT}) (s : seq {group gT}) :=
+ map section_repr (pairmap (@mkSec _) G s).
+
+End Sections.
+
+Section CompositionSeries.
+
+Variable gT : finGroupType.
+Local Notation gTg := {group gT}.
+Implicit Types (G : gTg) (s : seq gTg).
+
+Local Notation compo := [rel x y : {set gT} | maxnormal y x x].
+
+Definition comps G s := ((last G s) == 1%G) && compo.-series G s.
+
+Lemma compsP G s :
+ reflect (last G s = 1%G /\ path [rel x y : gTg | maxnormal y x x] G s)
+ (comps G s).
+Proof. by apply: (iffP andP) => [] [/eqP]. Qed.
+
+Lemma trivg_comps G s : comps G s -> (G :==: 1) = (s == [::]).
+Proof.
+case/andP=> ls cs; apply/eqP/eqP=> [G1 | s1]; last first.
+ by rewrite s1 /= in ls; apply/eqP.
+by case: s {ls} cs => //= H s /andP[/maxgroupp]; rewrite G1 /proper sub1G andbF.
+Qed.
+
+Lemma comps_cons G H s : comps G (H :: s) -> comps H s.
+Proof. by case/andP => /= ls /andP[_]; rewrite /comps ls. Qed.
+
+Lemma simple_compsP G s : comps G s -> reflect (s = [:: 1%G]) (simple G).
+Proof.
+move=> cs; apply: (iffP idP) => [|s1]; last first.
+ by rewrite s1 /comps eqxx /= andbT -simple_maxnormal in cs.
+case: s cs => [/trivg_comps/eqP-> | H s]; first by case/simpleP; rewrite eqxx.
+rewrite [comps _ _]andbCA /= => /andP[/maxgroupp maxH /trivg_comps/esym nil_s].
+rewrite simple_maxnormal => /maxgroupP[_ simG].
+have H1: H = 1%G by apply/val_inj/simG; rewrite // sub1G.
+by move: nil_s; rewrite H1 eqxx => /eqP->.
+Qed.
+
+Lemma exists_comps (G : gTg) : exists s, comps G s.
+Proof.
+elim: {G} #|G| {1 3}G (leqnn #|G|) => [G | n IHn G cG].
+ by rewrite leqNgt cardG_gt0.
+have [sG | nsG] := boolP (simple G).
+ by exists [:: 1%G]; rewrite /comps eqxx /= -simple_maxnormal andbT.
+have [-> | ntG] := eqVneq G 1%G; first by exists [::]; rewrite /comps eqxx.
+have [N maxN] := ex_maxnormal_ntrivg ntG.
+have [|s /andP[ls cs]] := IHn N.
+ by rewrite -ltnS (leq_trans _ cG) // proper_card // (maxnormal_proper maxN).
+by exists (N :: s); exact/and3P.
+Qed.
+
+(******************************************************************************)
+(* The factors associated to two composition series of the same group are *)
+(* the same up to isomorphism and permutation *)
+(******************************************************************************)
+
+Lemma JordanHolderUniqueness (G : gTg) (s1 s2 : seq gTg) :
+ comps G s1 -> comps G s2 -> perm_eq (mkfactors G s1) (mkfactors G s2).
+Proof.
+elim: {G}#|G| {-2}G (leqnn #|G|) => [|n Hi] G cG in s1 s2 * => cs1 cs2.
+ by rewrite leqNgt cardG_gt0 in cG.
+have [G1 | ntG] := boolP (G :==: 1).
+ have -> : s1 = [::] by apply/eqP; rewrite -(trivg_comps cs1).
+ have -> : s2 = [::] by apply/eqP; rewrite -(trivg_comps cs2).
+ by rewrite /= perm_eq_refl.
+have [sG | nsG] := boolP (simple G).
+ by rewrite (simple_compsP cs1 sG) (simple_compsP cs2 sG) perm_eq_refl.
+case es1: s1 cs1 => [|N1 st1] cs1.
+ by move: (trivg_comps cs1); rewrite eqxx; move/negP:ntG.
+case es2: s2 cs2 => [|N2 st2] cs2 {s1 es1}.
+ by move: (trivg_comps cs2); rewrite eqxx; move/negP:ntG.
+case/andP: cs1 => /= lst1; case/andP=> maxN_1 pst1.
+case/andP: cs2 => /= lst2; case/andP=> maxN_2 pst2.
+have cN1 : #|N1| <= n.
+ by rewrite -ltnS (leq_trans _ cG) ?proper_card ?(maxnormal_proper maxN_1).
+have cN2 : #|N2| <= n.
+ by rewrite -ltnS (leq_trans _ cG) ?proper_card ?(maxnormal_proper maxN_2).
+case: (N1 =P N2) {s2 es2} => [eN12 |].
+ by rewrite eN12 /= perm_cons Hi // /comps ?lst2 //= -eN12 lst1.
+move/eqP; rewrite -val_eqE /=; move/eqP=> neN12.
+have nN1G : N1 <| G by apply: maxnormal_normal.
+have nN2G : N2 <| G by apply: maxnormal_normal.
+pose N := (N1 :&: N2)%G.
+have nNG : N <| G.
+ by rewrite /normal subIset ?(normal_sub nN1G) //= normsI ?normal_norm.
+have iso1 : (G / N1)%G \isog (N2 / N)%G.
+ rewrite isog_sym /= -(maxnormalM maxN_1 maxN_2) //.
+ rewrite (@normC _ N1 N2) ?(subset_trans (normal_sub nN1G)) ?normal_norm //.
+ by rewrite weak_second_isog ?(subset_trans (normal_sub nN2G)) ?normal_norm.
+have iso2 : (G / N2)%G \isog (N1 / N)%G.
+ rewrite isog_sym /= -(maxnormalM maxN_1 maxN_2) // setIC.
+ by rewrite weak_second_isog ?(subset_trans (normal_sub nN1G)) ?normal_norm.
+have [sN /andP[lsN csN]] := exists_comps N.
+have i1 : perm_eq (mksrepr G N1 :: mkfactors N1 st1)
+ [:: mksrepr G N1, mksrepr N1 N & mkfactors N sN].
+ rewrite perm_cons -[mksrepr _ _ :: _]/(mkfactors N1 [:: N & sN]).
+ apply: Hi=> //; rewrite /comps ?lst1 //= lsN csN andbT /=.
+ rewrite -quotient_simple.
+ by rewrite -(isog_simple iso2) quotient_simple.
+ by rewrite (normalS (subsetIl N1 N2) (normal_sub nN1G)).
+have i2 : perm_eq (mksrepr G N2 :: mkfactors N2 st2)
+ [:: mksrepr G N2, mksrepr N2 N & mkfactors N sN].
+ rewrite perm_cons -[mksrepr _ _ :: _]/(mkfactors N2 [:: N & sN]).
+ apply: Hi=> //; rewrite /comps ?lst2 //= lsN csN andbT /=.
+ rewrite -quotient_simple.
+ by rewrite -(isog_simple iso1) quotient_simple.
+ by rewrite (normalS (subsetIr N1 N2) (normal_sub nN2G)).
+pose fG1 := [:: mksrepr G N1, mksrepr N1 N & mkfactors N sN].
+pose fG2 := [:: mksrepr G N2, mksrepr N2 N & mkfactors N sN].
+have i3 : perm_eq fG1 fG2.
+ rewrite (@perm_catCA _ [::_] [::_]) /mksrepr.
+ rewrite (@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso1).
+ rewrite -(@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso2).
+ exact: perm_eq_refl.
+apply: (perm_eq_trans i1); apply: (perm_eq_trans i3); rewrite perm_eq_sym.
+apply: perm_eq_trans i2; exact: perm_eq_refl.
+Qed.
+
+End CompositionSeries.
+
+(******************************************************************************)
+(* Helper lemmas for group actions. *)
+(******************************************************************************)
+
+Section MoreGroupAction.
+
+Variables (aT rT : finGroupType).
+Variables (A : {group aT}) (D : {group rT}).
+Variable to : groupAction A D.
+
+Lemma gactsP (G : {set rT}) : reflect {acts A, on G | to} [acts A, on G | to].
+Proof.
+apply: (iffP idP) => [nGA x|nGA]; first exact: acts_act.
+apply/subsetP=> a Aa; rewrite !inE; rewrite Aa.
+by apply/subsetP=> x; rewrite inE nGA.
+Qed.
+
+Lemma gactsM (N1 N2 : {set rT}) :
+ N1 \subset D -> N2 \subset D ->
+ [acts A, on N1 | to] -> [acts A, on N2 | to] -> [acts A, on N1 * N2 | to].
+Proof.
+move=> sN1D sN2D aAN1 aAN2; apply/gactsP=> x Ax y.
+apply/idP/idP; case/mulsgP=> y1 y2 N1y1 N2y2 e.
+ move: (actKin to Ax y); rewrite e; move<-.
+ rewrite gactM ?groupV ?(subsetP sN1D y1) ?(subsetP sN2D) //.
+ by apply: mem_mulg; rewrite ?(gactsP _ aAN1) ?(gactsP _ aAN2) // groupV.
+rewrite e gactM // ?(subsetP sN1D y1) ?(subsetP sN2D) //.
+by apply: mem_mulg; rewrite ?(gactsP _ aAN1) // ?(gactsP _ aAN2).
+Qed.
+
+Lemma gactsI (N1 N2 : {set rT}) :
+ [acts A, on N1 | to] -> [acts A, on N2 | to] -> [acts A, on N1 :&: N2 | to].
+Proof.
+move=> aAN1 aAN2.
+apply/subsetP=> x Ax; rewrite !inE Ax /=; apply/subsetP=> y Ny; rewrite inE.
+case/setIP: Ny=> N1y N2y; rewrite inE ?astabs_act ?N1y ?N2y //.
+- by move/subsetP: aAN2; move/(_ x Ax).
+- by move/subsetP: aAN1; move/(_ x Ax).
+Qed.
+
+Lemma gastabsP (S : {set rT}) (a : aT) :
+ a \in A -> reflect (forall x, (to x a \in S) = (x \in S)) (a \in 'N(S | to)).
+Proof.
+move=> Aa; apply: (iffP idP) => [nSa x|nSa]; first exact: astabs_act.
+by rewrite !inE Aa; apply/subsetP=> x; rewrite inE nSa.
+Qed.
+
+End MoreGroupAction.
+
+(******************************************************************************)
+(* Helper lemmas for quotient actions. *)
+(******************************************************************************)
+
+Section MoreQuotientAction.
+
+Variables (aT rT : finGroupType).
+Variables (A : {group aT})(D : {group rT}).
+Variable to : groupAction A D.
+
+Lemma qact_dom_doms (H : {group rT}) : H \subset D -> qact_dom to H \subset A.
+Proof.
+by move=> sHD; apply/subsetP=> x; rewrite qact_domE // inE; case/andP.
+Qed.
+
+Lemma acts_qact_doms (H : {group rT}) :
+ H \subset D -> [acts A, on H | to] -> qact_dom to H :=: A.
+Proof.
+move=> sHD aH; apply/eqP; rewrite eqEsubset; apply/andP.
+split; first exact: qact_dom_doms.
+apply/subsetP=> x Ax; rewrite qact_domE //; apply/gastabsP=> //.
+by move/gactsP: aH; move/(_ x Ax).
+Qed.
+
+Lemma qacts_cosetpre (H : {group rT}) (K' : {group coset_of H}) :
+ H \subset D -> [acts A, on H | to] ->
+ [acts qact_dom to H, on K' | to / H] ->
+ [acts A, on coset H @*^-1 K' | to].
+Proof.
+move=> sHD aH aK'; apply/subsetP=> x Ax; move: (Ax) (subsetP aK').
+rewrite -{1}(acts_qact_doms sHD aH) => qdx; move/(_ x qdx) => nx.
+rewrite !inE Ax; apply/subsetP=> y; case/morphpreP=> Ny /= K'Hy; rewrite inE.
+apply/morphpreP; split; first by rewrite acts_qact_dom_norm.
+by move/gastabsP: nx; move/(_ qdx (coset H y)); rewrite K'Hy qactE.
+Qed.
+
+Lemma qacts_coset (H K : {group rT}) :
+ H \subset D -> [acts A, on K | to] ->
+ [acts qact_dom to H, on (coset H) @* K | to / H].
+Proof.
+move=> sHD aK.
+apply/subsetP=> x qdx; rewrite inE qdx inE; apply/subsetP=> y.
+case/morphimP=> z Nz Kz /= e; rewrite e inE qactE // mem_imset // inE.
+move/gactsP: aK; move/(_ x (subsetP (qact_dom_doms sHD) _ qdx) z); rewrite Kz.
+move->; move/acts_act: (acts_qact_dom to H); move/(_ x qdx z).
+by rewrite Nz andbT.
+Qed.
+
+End MoreQuotientAction.
+
+Section StableCompositionSeries.
+
+Variables (aT rT : finGroupType).
+Variables (D : {group rT})(A : {group aT}).
+Variable to : groupAction A D.
+
+Definition maxainv (B C : {set rT}) :=
+ [max C of H |
+ [&& (H <| B), ~~ (B \subset H) & [acts A, on H | to]]].
+
+Section MaxAinvProps.
+
+Variables K N : {group rT}.
+
+Lemma maxainv_norm : maxainv K N -> N <| K.
+Proof. by move/maxgroupp; case/andP. Qed.
+
+Lemma maxainv_proper : maxainv K N -> N \proper K.
+Proof.
+by move/maxgroupp; case/andP; rewrite properE; move/normal_sub->; case/andP.
+Qed.
+
+Lemma maxainv_sub : maxainv K N -> N \subset K.
+Proof. move=> h; apply: proper_sub; exact: maxainv_proper. Qed.
+
+Lemma maxainv_ainvar : maxainv K N -> A \subset 'N(N | to).
+Proof. by move/maxgroupp; case/and3P. Qed.
+
+Lemma maxainvS : maxainv K N -> N \subset K.
+Proof. by move=> pNN; rewrite proper_sub // maxainv_proper. Qed.
+
+Lemma maxainv_exists : K :!=: 1 -> {N : {group rT} | maxainv K N}.
+Proof.
+move=> nt; apply: ex_maxgroup. exists [1 rT]%G.
+rewrite /= normal1 subG1 nt /=.
+apply/subsetP=> a Da; rewrite !inE Da /= sub1set !inE.
+by rewrite /= -actmE // morph1 eqxx.
+Qed.
+
+End MaxAinvProps.
+
+Lemma maxainvM (G H K : {group rT}) :
+ H \subset D -> K \subset D -> maxainv G H -> maxainv G K ->
+ H :<>: K -> H * K = G.
+Proof.
+move: H K => N1 N2 sN1D sN2D pmN1 pmN2 neN12.
+have cN12 : commute N1 N2.
+ apply: normC; apply: (subset_trans (maxainv_sub pmN1)).
+ by rewrite normal_norm ?maxainv_norm.
+wlog nsN21 : G N1 N2 sN1D sN2D pmN1 pmN2 neN12 cN12/ ~~(N1 \subset N2).
+ move/eqP: (neN12); rewrite eqEsubset negb_and; case/orP=> ns; first by apply.
+ by rewrite cN12; apply=> //; apply: sym_not_eq.
+have nP : N1 * N2 <| G by rewrite normalM ?maxainv_norm.
+have sN2P : N2 \subset N1 * N2 by rewrite mulg_subr ?group1.
+case/maxgroupP: (pmN1); case/andP=> nN1G pN1G mN1.
+case/maxgroupP: (pmN2); case/andP=> nN2G pN2G mN2.
+case/andP: pN1G=> nsGN1 ha1; case/andP: pN2G=> nsGN2 ha2.
+case e : (G \subset N1 * N2).
+ by apply/eqP; rewrite eqEsubset e mulG_subG !normal_sub.
+have: N1 <*> N2 = N2 by apply: mN2; rewrite /= ?comm_joingE // nP e /= gactsM.
+by rewrite comm_joingE // => h; move: nsN21; rewrite -h mulg_subl.
+Qed.
+
+Definition asimple (K : {set rT}) := maxainv K 1.
+
+Implicit Types (H K : {group rT}) (s : seq {group rT}).
+
+Lemma asimpleP K :
+ reflect [/\ K :!=: 1
+ & forall H, H <| K -> [acts A, on H | to] -> H :=: 1 \/ H :=: K]
+ (asimple K).
+Proof.
+apply: (iffP idP).
+ case/maxgroupP; rewrite normal1 /=; case/andP=> nsK1 aK H1.
+ rewrite eqEsubset negb_and nsK1 /=; split => // H nHK ha.
+ case eHK : (H :==: K); first by right; apply/eqP.
+ left; apply: H1; rewrite ?sub1G // nHK; move/negbT: eHK.
+ by rewrite eqEsubset negb_and normal_sub //=; move->.
+case=> ntK h; apply/maxgroupP; split.
+ move: ntK; rewrite eqEsubset sub1G andbT normal1; move->.
+ apply/subsetP=> a Da; rewrite !inE Da /= sub1set !inE.
+ by rewrite /= -actmE // morph1 eqxx.
+move=> H /andP[nHK /andP[nsKH ha]] _.
+case: (h _ nHK ha)=> // /eqP; rewrite eqEsubset.
+by rewrite (negbTE nsKH) andbF.
+Qed.
+
+Definition acomps K s :=
+ ((last K s) == 1%G) && path [rel x y : {group rT} | maxainv x y] K s.
+
+Lemma acompsP K s :
+ reflect (last K s = 1%G /\ path [rel x y : {group rT} | maxainv x y] K s)
+ (acomps K s).
+Proof. by apply: (iffP andP); case; move/eqP. Qed.
+
+Lemma trivg_acomps K s : acomps K s -> (K :==: 1) = (s == [::]).
+Proof.
+case/andP=> ls cs; apply/eqP/eqP; last first.
+ by move=> se; rewrite se /= in ls; apply/eqP.
+move=> G1; case: s ls cs => // H s _ /=; case/andP; case/maxgroupP.
+by rewrite G1 sub1G andbF.
+Qed.
+
+Lemma acomps_cons K H s : acomps K (H :: s) -> acomps H s.
+Proof. by case/andP => /= ls; case/andP=> _ p; rewrite /acomps ls. Qed.
+
+Lemma asimple_acompsP K s : acomps K s -> reflect (s = [:: 1%G]) (asimple K).
+Proof.
+move=> cs; apply: (iffP idP); last first.
+ by move=> se; move: cs; rewrite se /=; case/andP=> /=; rewrite andbT.
+case: s cs.
+ by rewrite /acomps /= andbT; move/eqP->; case/asimpleP; rewrite eqxx.
+move=> H s cs sG; apply/eqP.
+rewrite eqseq_cons -(trivg_acomps (acomps_cons cs)) andbC andbb.
+case/acompsP: cs => /= ls; case/andP=> mH ps.
+case/maxgroupP: sG; case/and3P => _ ntG _ ->; rewrite ?sub1G //.
+rewrite (maxainv_norm mH); case/andP: (maxainv_proper mH)=> _ ->.
+exact: (maxainv_ainvar mH).
+Qed.
+
+Lemma exists_acomps K : exists s, acomps K s.
+Proof.
+elim: {K} #|K| {1 3}K (leqnn #|K|) => [K | n Hi K cK].
+ by rewrite leqNgt cardG_gt0.
+case/orP: (orbN (asimple K)) => [sK | nsK].
+ by exists [:: (1%G : {group rT})]; rewrite /acomps eqxx /= andbT.
+case/orP: (orbN (K :==: 1))=> [tK | ntK].
+ by exists (Nil _); rewrite /acomps /= andbT.
+case: (maxainv_exists ntK)=> N pmN.
+have cN: #|N| <= n.
+ by rewrite -ltnS (leq_trans _ cK) // proper_card // (maxainv_proper pmN).
+case: (Hi _ cN)=> s; case/andP=> lasts ps; exists [:: N & s]; rewrite /acomps.
+by rewrite last_cons lasts /= pmN.
+Qed.
+
+End StableCompositionSeries.
+
+Arguments Scope maxainv
+ [_ _ Group_scope Group_scope groupAction_scope group_scope group_scope].
+Arguments Scope asimple
+ [_ _ Group_scope Group_scope groupAction_scope group_scope].
+
+Section StrongJordanHolder.
+
+Section AuxiliaryLemmas.
+
+Variables aT rT : finGroupType.
+Variables (A : {group aT}) (D : {group rT}) (to : groupAction A D).
+
+Lemma maxainv_asimple_quo (G H : {group rT}) :
+ H \subset D -> maxainv to G H -> asimple (to / H) (G / H).
+Proof.
+move=> sHD /maxgroupP[/and3P[nHG pHG aH] Hmax].
+apply/asimpleP; split; first by rewrite -subG1 quotient_sub1 ?normal_norm.
+move=> K' nK'Q aK'.
+have: (K' \proper (G / H)) || (G / H == K').
+ by rewrite properE eqEsubset andbC (normal_sub nK'Q) !andbT orbC orbN.
+case/orP=> [ pHQ | eQH]; last by right; apply sym_eq; apply/eqP.
+left; pose K := ((coset H) @*^-1 K')%G.
+have eK'I : K' \subset (coset H) @* 'N(H).
+ by rewrite (subset_trans (normal_sub nK'Q)) ?morphimS ?normal_norm.
+have eKK' : K' :=: K / H by rewrite /(K / H) morphpreK //=.
+suff eKH : K :=: H by rewrite -trivg_quotient eKK' eKH.
+have sHK : H \subset K by rewrite -ker_coset kerE morphpreS // sub1set group1.
+apply: Hmax => //; apply/and3P; split; last exact: qacts_cosetpre.
+ by rewrite -(quotientGK nHG) cosetpre_normal.
+by move: (proper_subn pHQ); rewrite sub_morphim_pre ?normal_norm.
+Qed.
+
+
+Lemma asimple_quo_maxainv (G H : {group rT}) :
+ H \subset D -> G \subset D -> [acts A, on G | to] -> [acts A, on H | to] ->
+ H <| G -> asimple (to / H) (G / H) ->
+ maxainv to G H.
+Proof.
+move=> sHD sGD aG aH nHG /asimpleP[ntQ maxQ]; apply/maxgroupP; split.
+ by rewrite nHG -quotient_sub1 ?normal_norm // subG1 ntQ.
+move=> K /and3P[nKG nsGK aK] sHK.
+pose K' := (K / H)%G.
+have K'dQ : K' <| (G / H)%G by apply: morphim_normal.
+have nKH : H <| K by rewrite (normalS _ _ nHG) // normal_sub.
+have: K' :=: 1%G \/ K' :=: (G / H).
+ apply: (maxQ K' K'dQ) => /=.
+ apply/subsetP=> x Adx. rewrite inE Adx /= inE. apply/subsetP=> y.
+ rewrite quotientE; case/morphimP=> z Nz Kz ->; rewrite /= !inE qactE //.
+ have ntoyx : to z x \in 'N(H) by rewrite (acts_qact_dom_norm Adx).
+ apply/morphimP; exists (to z x) => //.
+ suff h: qact_dom to H \subset A.
+ by rewrite astabs_act // (subsetP aK) //; apply: (subsetP h).
+ by apply/subsetP=> t; rewrite qact_domE // inE; case/ andP.
+case; last first.
+ move/quotient_injG; rewrite !inE /=; move/(_ nKH nHG)=> c; move: nsGK.
+ by rewrite c subxx.
+rewrite /= -trivg_quotient; move=> tK'; apply:(congr1 (@gval _)); move: tK'.
+by apply: (@quotient_injG _ H); rewrite ?inE /= ?normal_refl.
+Qed.
+
+Lemma asimpleI (N1 N2 : {group rT}) :
+ N2 \subset 'N(N1) -> N1 \subset D ->
+ [acts A, on N1 | to] -> [acts A, on N2 | to] ->
+ asimple (to / N1) (N2 / N1) ->
+ asimple (to / (N2 :&: N1)) (N2 / (N2 :&: N1)).
+Proof.
+move=> nN21 sN1D aN1 aN2 /asimpleP[ntQ1 max1].
+have [f1 [f1e f1ker f1pre f1im]] := restrmP (coset_morphism N1) nN21.
+have hf2' : N2 \subset 'N(N2 :&: N1) by apply: normsI => //; rewrite normG.
+have hf2'' : 'ker (coset (N2 :&: N1)) \subset 'ker f1.
+ by rewrite f1ker !ker_coset.
+pose f2 := factm_morphism hf2'' hf2'.
+apply/asimpleP; split.
+ rewrite /= setIC; apply/negP; move: (second_isog nN21); move/isog_eq1->.
+ by apply/negP.
+move=> H nHQ2 aH; pose K := f2 @* H.
+have nKQ1 : K <| N2 / N1.
+ rewrite (_ : N2 / N1 = f2 @* (N2 / (N2 :&: N1))) ?morphim_normal //.
+ by rewrite morphim_factm f1im.
+have sqA : qact_dom to N1 \subset A.
+ by apply/subsetP=> t; rewrite qact_domE // inE; case/andP.
+have nNN2 : (N2 :&: N1) <| N2.
+ rewrite /normal subsetIl; apply: normsI => //; exact: normG.
+have aKQ1 : [acts qact_dom to N1, on K | to / N1].
+ pose H':= coset (N2 :&: N1)@*^-1 H.
+ have eHH' : H :=: H' / (N2 :&: N1) by rewrite cosetpreK.
+ have -> : K :=: f1 @* H' by rewrite /K eHH' morphim_factm.
+ have sH'N2 : H' \subset N2.
+ rewrite /H' eHH' quotientGK ?normal_cosetpre //=.
+ by rewrite sub_cosetpre_quo ?normal_sub.
+ have -> : f1 @* H' = coset N1 @* H' by rewrite f1im //=.
+ apply: qacts_coset => //; apply: qacts_cosetpre => //; last exact: gactsI.
+ by apply: (subset_trans (subsetIr _ _)).
+have injf2 : 'injm f2.
+ by rewrite ker_factm f1ker /= ker_coset /= subG1 /= -quotientE trivg_quotient.
+have iHK : H \isog K.
+ apply/isogP; pose f3 := restrm_morphism (normal_sub nHQ2) f2.
+ by exists f3; rewrite 1?injm_restrm // morphim_restrm setIid.
+case: (max1 _ nKQ1 aKQ1).
+ by move/eqP; rewrite -(isog_eq1 iHK); move/eqP->; left.
+move=> he /=; right; apply/eqP; rewrite eqEcard normal_sub //=.
+move: (second_isog nN21); rewrite setIC; move/card_isog->; rewrite -he.
+by move/card_isog: iHK=> <-; rewrite leqnn.
+Qed.
+
+End AuxiliaryLemmas.
+
+Variables (aT rT : finGroupType).
+Variables (A : {group aT}) (D : {group rT}) (to : groupAction A D).
+
+(******************************************************************************)
+(* The factors associated to two A-stable composition series of the same *)
+(* group are the same up to isomorphism and permutation *)
+(******************************************************************************)
+
+Lemma StrongJordanHolderUniqueness (G : {group rT}) (s1 s2 : seq {group rT}) :
+ G \subset D -> acomps to G s1 -> acomps to G s2 ->
+ perm_eq (mkfactors G s1) (mkfactors G s2).
+Proof.
+elim: {G} #|G| {-2}G (leqnn #|G|) => [|n Hi] G cG in s1 s2 * => hsD cs1 cs2.
+ by rewrite leqNgt cardG_gt0 in cG.
+case/orP: (orbN (G :==: 1)) => [tG | ntG].
+ have -> : s1 = [::] by apply/eqP; rewrite -(trivg_acomps cs1).
+ have -> : s2 = [::] by apply/eqP; rewrite -(trivg_acomps cs2).
+ by rewrite /= perm_eq_refl.
+case/orP: (orbN (asimple to G))=> [sG | nsG].
+ have -> : s1 = [:: 1%G ] by apply/(asimple_acompsP cs1).
+ have -> : s2 = [:: 1%G ] by apply/(asimple_acompsP cs2).
+ by rewrite /= perm_eq_refl.
+case es1: s1 cs1 => [|N1 st1] cs1.
+ by move: (trivg_comps cs1); rewrite eqxx; move/negP:ntG.
+case es2: s2 cs2 => [|N2 st2] cs2 {s1 es1}.
+ by move: (trivg_comps cs2); rewrite eqxx; move/negP:ntG.
+case/andP: cs1 => /= lst1; case/andP=> maxN_1 pst1.
+case/andP: cs2 => /= lst2; case/andP=> maxN_2 pst2.
+have sN1D : N1 \subset D.
+ by apply: subset_trans hsD; apply: maxainv_sub maxN_1.
+have sN2D : N2 \subset D.
+ by apply: subset_trans hsD; apply: maxainv_sub maxN_2.
+have cN1 : #|N1| <= n.
+ by rewrite -ltnS (leq_trans _ cG) ?proper_card ?(maxainv_proper maxN_1).
+have cN2 : #|N2| <= n.
+ by rewrite -ltnS (leq_trans _ cG) ?proper_card ?(maxainv_proper maxN_2).
+case: (N1 =P N2) {s2 es2} => [eN12 |].
+ by rewrite eN12 /= perm_cons Hi // /acomps ?lst2 //= -eN12 lst1.
+move/eqP; rewrite -val_eqE /=; move/eqP=> neN12.
+have nN1G : N1 <| G by apply: (maxainv_norm maxN_1).
+have nN2G : N2 <| G by apply: (maxainv_norm maxN_2).
+pose N := (N1 :&: N2)%G.
+have nNG : N <| G.
+ by rewrite /normal subIset ?(normal_sub nN1G) //= normsI ?normal_norm.
+have iso1 : (G / N1)%G \isog (N2 / N)%G.
+ rewrite isog_sym /= -(maxainvM _ _ maxN_1 maxN_2) //.
+ rewrite (@normC _ N1 N2) ?(subset_trans (normal_sub nN1G)) ?normal_norm //.
+ by rewrite weak_second_isog ?(subset_trans (normal_sub nN2G)) ?normal_norm.
+have iso2 : (G / N2)%G \isog (N1 / N)%G.
+ rewrite isog_sym /= -(maxainvM _ _ maxN_1 maxN_2) // setIC.
+ by rewrite weak_second_isog ?(subset_trans (normal_sub nN1G)) ?normal_norm.
+case: (exists_acomps to N)=> sN; case/andP=> lsN csN.
+have aN1 : [acts A, on N1 | to].
+ by case/maxgroupP: maxN_1; case/and3P.
+have aN2 : [acts A, on N2 | to].
+ by case/maxgroupP: maxN_2; case/and3P.
+have nNN1 : N <| N1.
+ by apply: (normalS _ _ nNG); rewrite ?subsetIl ?normal_sub.
+have nNN2 : N <| N2.
+ by apply: (normalS _ _ nNG); rewrite ?subsetIr ?normal_sub.
+have aN : [ acts A, on N1 :&: N2 | to].
+ apply/subsetP=> x Ax; rewrite !inE Ax /=; apply/subsetP=> y Ny; rewrite inE.
+ case/setIP: Ny=> N1y N2y. rewrite inE ?astabs_act ?N1y ?N2y //.
+ by move/subsetP: aN2; move/(_ x Ax).
+ by move/subsetP: aN1; move/(_ x Ax).
+have i1 : perm_eq (mksrepr G N1 :: mkfactors N1 st1)
+ [:: mksrepr G N1, mksrepr N1 N & mkfactors N sN].
+ rewrite perm_cons -[mksrepr _ _ :: _]/(mkfactors N1 [:: N & sN]).
+ apply: Hi=> //; rewrite /acomps ?lst1 //= lsN csN andbT /=.
+ apply: asimple_quo_maxainv=> //; first by apply: subIset; rewrite sN1D.
+ apply: asimpleI => //.
+ apply: subset_trans (normal_norm nN2G); exact: normal_sub.
+ rewrite -quotientMidl (maxainvM _ _ maxN_2) //.
+ by apply: maxainv_asimple_quo.
+ by move=> e; apply: neN12.
+have i2 : perm_eq (mksrepr G N2 :: mkfactors N2 st2)
+ [:: mksrepr G N2, mksrepr N2 N & mkfactors N sN].
+ rewrite perm_cons -[mksrepr _ _ :: _]/(mkfactors N2 [:: N & sN]).
+ apply: Hi=> //; rewrite /acomps ?lst2 //= lsN csN andbT /=.
+ apply: asimple_quo_maxainv=> //; first by apply: subIset; rewrite sN1D.
+ have e : N1 :&: N2 :=: N2 :&: N1 by rewrite setIC.
+ rewrite (group_inj (setIC N1 N2)); apply: asimpleI => //.
+ apply: subset_trans (normal_norm nN1G); exact: normal_sub.
+ rewrite -quotientMidl (maxainvM _ _ maxN_1) //.
+ exact: maxainv_asimple_quo.
+pose fG1 := [:: mksrepr G N1, mksrepr N1 N & mkfactors N sN].
+pose fG2 := [:: mksrepr G N2, mksrepr N2 N & mkfactors N sN].
+have i3 : perm_eq fG1 fG2.
+ rewrite (@perm_catCA _ [::_] [::_]) /mksrepr.
+ rewrite (@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso1).
+ rewrite -(@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso2).
+ exact: perm_eq_refl.
+apply: (perm_eq_trans i1); apply: (perm_eq_trans i3); rewrite perm_eq_sym.
+apply: perm_eq_trans i2; exact: perm_eq_refl.
+Qed.
+
+End StrongJordanHolder.
+
+
+
+
+
diff --git a/mathcomp/solvable/maximal.v b/mathcomp/solvable/maximal.v
new file mode 100644
index 0000000..7b50dc9
--- /dev/null
+++ b/mathcomp/solvable/maximal.v
@@ -0,0 +1,1656 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice div fintype.
+Require Import finfun bigop finset prime binomial fingroup morphism perm.
+Require Import automorphism quotient action commutator gproduct gfunctor.
+Require Import ssralg finalg zmodp cyclic pgroup center gseries.
+Require Import nilpotent sylow abelian finmodule.
+
+(******************************************************************************)
+(* This file establishes basic properties of several important classes of *)
+(* maximal subgroups: maximal, max and min normal, simple, characteristically *)
+(* simple subgroups, the Frattini and Fitting subgroups, the Thompson *)
+(* critical subgroup, special and extra-special groups, and self-centralising *)
+(* normal (SCN) subgroups. In detail, we define: *)
+(* charsimple G == G is characteristically simple (it has no nontrivial *)
+(* characteristic subgroups, and is nontrivial) *)
+(* 'Phi(G) == the Frattini subgroup of G, i.e., the intersection of *)
+(* all its maximal proper subgroups. *)
+(* 'F(G) == the Fitting subgroup of G, i.e., the largest normal *)
+(* nilpotent subgroup of G (defined as the (direct) *)
+(* product of all the p-cores of G). *)
+(* critical C G == C is a critical subgroup of G: C is characteristic *)
+(* (but not functorial) in G, the center of C contains *)
+(* both its Frattini subgroup and the commutator [G, C], *)
+(* and is equal to the centraliser of C in G. The *)
+(* Thompson_critical theorem provides critical subgroups *)
+(* for p-groups; we also show that in this case the *)
+(* centraliser of C in Aut G is a p-group as well. *)
+(* special G == G is a special group: its center, Frattini, and *)
+(* derived sugroups coincide (we follow Aschbacher in *)
+(* not considering nontrivial elementary abelian groups *)
+(* as special); we show that a p-group factors under *)
+(* coprime action into special groups (Aschbacher 24.7). *)
+(* extraspecial G == G is a special group whose center has prime order *)
+(* (hence G is non-abelian). *)
+(* 'SCN(G) == the set of self-centralising normal abelian subgroups *)
+(* of G (the A <| G such that 'C_G(A) = A). *)
+(* 'SCN_n(G) == the subset of 'SCN(G) containing all groups with rank *)
+(* at least n (i.e., A \in 'SCN(G) and 'm(A) >= n). *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section Defs.
+
+Variable gT : finGroupType.
+Implicit Types (A B D : {set gT}) (G : {group gT}).
+
+Definition charsimple A := [min A of G | G :!=: 1 & G \char A].
+
+Definition Frattini A := \bigcap_(G : {group gT} | maximal_eq G A) G.
+
+Canonical Frattini_group A : {group gT} := Eval hnf in [group of Frattini A].
+
+Definition Fitting A := \big[dprod/1]_(p <- primes #|A|) 'O_p(A).
+
+Lemma Fitting_group_set G : group_set (Fitting G).
+Proof.
+suffices [F ->]: exists F : {group gT}, Fitting G = F by exact: groupP.
+rewrite /Fitting; elim: primes (primes_uniq #|G|) => [_|p r IHr] /=.
+ by exists [1 gT]%G; rewrite big_nil.
+case/andP=> rp /IHr[F defF]; rewrite big_cons defF.
+suffices{IHr} /and3P[p'F sFG nFG]: p^'.-group F && (F <| G).
+ have nFGp: 'O_p(G) \subset 'N(F) := subset_trans (pcore_sub p G) nFG.
+ have pGp: p.-group('O_p(G)) := pcore_pgroup p G.
+ have{pGp} tiGpF: 'O_p(G) :&: F = 1 by rewrite coprime_TIg ?(pnat_coprime pGp).
+ exists ('O_p(G) <*> F)%G; rewrite dprodEY // (sameP commG1P trivgP) -tiGpF.
+ by rewrite subsetI commg_subl commg_subr (subset_trans sFG) // gFnorm.
+move/bigdprodWY: defF => <- {F}; elim: r rp => [_|q r IHr] /=.
+ by rewrite big_nil gen0 pgroup1 normal1.
+rewrite inE eq_sym big_cons -joingE -joing_idr => /norP[qp /IHr {IHr}].
+set F := <<_>> => /andP[p'F nsFG]; rewrite norm_joinEl /= -/F; last first.
+ exact: subset_trans (pcore_sub q G) (normal_norm nsFG).
+by rewrite pgroupM p'F normalM ?pcore_normal //= (pi_pgroup (pcore_pgroup q G)).
+Qed.
+
+Canonical Fitting_group G := group (Fitting_group_set G).
+
+Definition critical A B :=
+ [/\ A \char B,
+ Frattini A \subset 'Z(A),
+ [~: B, A] \subset 'Z(A)
+ & 'C_B(A) = 'Z(A)].
+
+Definition special A := Frattini A = 'Z(A) /\ A^`(1) = 'Z(A).
+
+Definition extraspecial A := special A /\ prime #|'Z(A)|.
+
+Definition SCN B := [set A : {group gT} | A <| B & 'C_B(A) == A].
+
+Definition SCN_at n B := [set A in SCN B | n <= 'r(A)].
+
+End Defs.
+
+Arguments Scope charsimple [_ group_scope].
+Arguments Scope Frattini [_ group_scope].
+Arguments Scope Fitting [_ group_scope].
+Arguments Scope critical [_ group_scope group_scope].
+Arguments Scope special [_ group_scope].
+Arguments Scope extraspecial [_ group_scope].
+Arguments Scope SCN [_ group_scope].
+Arguments Scope SCN_at [_ nat_scope group_scope].
+
+Prenex Implicits maximal simple charsimple critical special extraspecial.
+
+Notation "''Phi' ( A )" := (Frattini A)
+ (at level 8, format "''Phi' ( A )") : group_scope.
+Notation "''Phi' ( G )" := (Frattini_group G) : Group_scope.
+
+Notation "''F' ( G )" := (Fitting G)
+ (at level 8, format "''F' ( G )") : group_scope.
+Notation "''F' ( G )" := (Fitting_group G) : Group_scope.
+
+Notation "''SCN' ( B )" := (SCN B)
+ (at level 8, format "''SCN' ( B )") : group_scope.
+Notation "''SCN_' n ( B )" := (SCN_at n B)
+ (at level 8, n at level 2, format "''SCN_' n ( B )") : group_scope.
+
+Section PMax.
+
+Variables (gT : finGroupType) (p : nat) (P M : {group gT}).
+Hypothesis pP : p.-group P.
+
+Lemma p_maximal_normal : maximal M P -> M <| P.
+Proof.
+case/maxgroupP=> /andP[sMP sPM] maxM; rewrite /normal sMP.
+have:= subsetIl P 'N(M); rewrite subEproper.
+case/predU1P=> [/setIidPl-> // | /maxM/= SNM]; case/negP: sPM.
+rewrite (nilpotent_sub_norm (pgroup_nil pP) sMP) //.
+by rewrite SNM // subsetI sMP normG.
+Qed.
+
+Lemma p_maximal_index : maximal M P -> #|P : M| = p.
+Proof.
+move=> maxM; have nM := p_maximal_normal maxM.
+rewrite -card_quotient ?normal_norm //.
+rewrite -(quotient_maximal _ nM) ?normal_refl // trivg_quotient in maxM.
+case/maxgroupP: maxM; rewrite properEneq eq_sym sub1G andbT /=.
+case/(pgroup_pdiv (quotient_pgroup M pP)) => p_pr /Cauchy[] // xq.
+rewrite /order -cycle_subG subEproper => /predU1P[-> // | sxPq oxq_p _].
+by move/(_ _ sxPq (sub1G _)) => xq1; rewrite -oxq_p xq1 cards1 in p_pr.
+Qed.
+
+Lemma p_index_maximal : M \subset P -> prime #|P : M| -> maximal M P.
+Proof.
+move=> sMP /primeP[lt1PM pr_PM].
+apply/maxgroupP; rewrite properEcard sMP -(Lagrange sMP).
+rewrite -{1}(muln1 #|M|) ltn_pmul2l //; split=> // H sHP sMH.
+apply/eqP; rewrite eq_sym eqEcard sMH.
+case/orP: (pr_PM _ (indexSg sMH (proper_sub sHP))) => /eqP iM.
+ by rewrite -(Lagrange sMH) iM muln1 /=.
+by have:= proper_card sHP; rewrite -(Lagrange sMH) iM Lagrange ?ltnn.
+Qed.
+
+End PMax.
+
+Section Frattini.
+
+Variables gT : finGroupType.
+Implicit Type G M : {group gT}.
+
+Lemma Phi_sub G : 'Phi(G) \subset G.
+Proof. by rewrite bigcap_inf // /maximal_eq eqxx. Qed.
+
+Lemma Phi_sub_max G M : maximal M G -> 'Phi(G) \subset M.
+Proof. by move=> maxM; rewrite bigcap_inf // /maximal_eq predU1r. Qed.
+
+Lemma Phi_proper G : G :!=: 1 -> 'Phi(G) \proper G.
+Proof.
+move/eqP; case/maximal_exists: (sub1G G) => [<- //| [M maxM _] _].
+exact: sub_proper_trans (Phi_sub_max maxM) (maxgroupp maxM).
+Qed.
+
+Lemma Phi_nongen G X : 'Phi(G) <*> X = G -> <<X>> = G.
+Proof.
+move=> defG; have: <<X>> \subset G by rewrite -{1}defG genS ?subsetUr.
+case/maximal_exists=> //= [[M maxM]]; rewrite gen_subG => sXM.
+case/andP: (maxgroupp maxM) => _ /negP[].
+by rewrite -defG gen_subG subUset Phi_sub_max.
+Qed.
+
+Lemma Frattini_continuous (rT : finGroupType) G (f : {morphism G >-> rT}) :
+ f @* 'Phi(G) \subset 'Phi(f @* G).
+Proof.
+apply/bigcapsP=> M maxM; rewrite sub_morphim_pre ?Phi_sub // bigcap_inf //.
+have {2}<-: f @*^-1 (f @* G) = G by rewrite morphimGK ?subsetIl.
+by rewrite morphpre_maximal_eq ?maxM //; case/maximal_eqP: maxM.
+Qed.
+
+End Frattini.
+
+Canonical Frattini_igFun := [igFun by Phi_sub & Frattini_continuous].
+Canonical Frattini_gFun := [gFun by Frattini_continuous].
+
+Section Frattini0.
+
+Variable gT : finGroupType.
+Implicit Types (rT : finGroupType) (D G : {group gT}).
+
+Lemma Phi_char G : 'Phi(G) \char G.
+Proof. exact: gFchar. Qed.
+
+Lemma Phi_normal G : 'Phi(G) <| G.
+Proof. exact: gFnormal. Qed.
+
+Lemma injm_Phi rT D G (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> f @* 'Phi(G) = 'Phi(f @* G).
+Proof. exact: injmF. Qed.
+
+Lemma isog_Phi rT G (H : {group rT}) : G \isog H -> 'Phi(G) \isog 'Phi(H).
+Proof. exact: gFisog. Qed.
+
+Lemma PhiJ G x : 'Phi(G :^ x) = 'Phi(G) :^ x.
+Proof.
+rewrite -{1}(setIid G) -(setIidPr (Phi_sub G)) -!morphim_conj.
+by rewrite injm_Phi ?injm_conj.
+Qed.
+
+End Frattini0.
+
+Section Frattini2.
+
+Variables gT : finGroupType.
+Implicit Type G : {group gT}.
+
+Lemma Phi_quotient_id G : 'Phi (G / 'Phi(G)) = 1.
+Proof.
+apply/trivgP; rewrite -cosetpreSK cosetpre1 /=; apply/bigcapsP=> M maxM.
+have nPhi := Phi_normal G; have nPhiM: 'Phi(G) <| M.
+ by apply: normalS nPhi; [exact: bigcap_inf | case/maximal_eqP: maxM].
+by rewrite sub_cosetpre_quo ?bigcap_inf // quotient_maximal_eq.
+Qed.
+
+Lemma Phi_quotient_cyclic G : cyclic (G / 'Phi(G)) -> cyclic G.
+Proof.
+case/cyclicP=> /= Px; case: (cosetP Px) => x nPx ->{Px} defG.
+apply/cyclicP; exists x; symmetry; apply: Phi_nongen.
+rewrite -joing_idr norm_joinEr -?quotientK ?cycle_subG //.
+by rewrite /quotient morphim_cycle //= -defG quotientGK ?Phi_normal.
+Qed.
+
+Variables (p : nat) (P : {group gT}).
+
+Lemma trivg_Phi : p.-group P -> ('Phi(P) == 1) = p.-abelem P.
+Proof.
+move=> pP; case: (eqsVneq P 1) => [P1 | ntP].
+ by rewrite P1 abelem1 -subG1 -P1 Phi_sub.
+have [p_pr _ _] := pgroup_pdiv pP ntP.
+apply/eqP/idP=> [trPhi | abP].
+ apply/abelemP=> //; split=> [|x Px].
+ apply/commG1P/trivgP; rewrite -trPhi.
+ apply/bigcapsP=> M /predU1P[-> | maxM]; first exact: der1_subG.
+ have /andP[_ nMP]: M <| P := p_maximal_normal pP maxM.
+ rewrite der1_min // cyclic_abelian // prime_cyclic // card_quotient //.
+ by rewrite (p_maximal_index pP).
+ apply/set1gP; rewrite -trPhi; apply/bigcapP=> M.
+ case/predU1P=> [-> | maxM]; first exact: groupX.
+ have /andP[_ nMP] := p_maximal_normal pP maxM.
+ have nMx : x \in 'N(M) by exact: subsetP Px.
+ apply: coset_idr; rewrite ?groupX ?morphX //=; apply/eqP.
+ rewrite -(p_maximal_index pP maxM) -card_quotient // -order_dvdn cardSg //=.
+ by rewrite cycle_subG mem_quotient.
+apply/trivgP/subsetP=> x Phi_x; rewrite -cycle_subG.
+have Px: x \in P by exact: (subsetP (Phi_sub P)).
+have sxP: <[x]> \subset P by rewrite cycle_subG.
+case/splitsP: (abelem_splits abP sxP) => K /complP[tiKx defP].
+have [-> | nt_x] := eqVneq x 1; first by rewrite cycle1.
+have oxp := abelem_order_p abP Px nt_x.
+rewrite /= -tiKx subsetI subxx cycle_subG.
+apply: (bigcapP Phi_x); apply/orP; right.
+apply: p_index_maximal; rewrite -?divgS -defP ?mulG_subr //.
+by rewrite (TI_cardMg tiKx) mulnK // [#|_|]oxp.
+Qed.
+
+End Frattini2.
+
+Section Frattini3.
+
+Variables (gT : finGroupType) (p : nat) (P : {group gT}).
+Hypothesis pP : p.-group P.
+
+Lemma Phi_quotient_abelem : p.-abelem (P / 'Phi(P)).
+Proof. by rewrite -trivg_Phi ?morphim_pgroup //= Phi_quotient_id. Qed.
+
+Lemma Phi_joing : 'Phi(P) = P^`(1) <*> 'Mho^1(P).
+Proof.
+have [sPhiP nPhiP] := andP (Phi_normal P).
+apply/eqP; rewrite eqEsubset join_subG.
+case: (eqsVneq P 1) => [-> | ntP] in sPhiP *.
+ by rewrite /= (trivgP sPhiP) sub1G der_subS Mho_sub.
+have [p_pr _ _] := pgroup_pdiv pP ntP.
+have [abP x1P] := abelemP p_pr Phi_quotient_abelem.
+apply/andP; split.
+ have nMP: P \subset 'N(P^`(1) <*> 'Mho^1(P)) by rewrite normsY // !gFnorm.
+ rewrite -quotient_sub1 ?(subset_trans sPhiP) //=.
+ suffices <-: 'Phi(P / (P^`(1) <*> 'Mho^1(P))) = 1 by exact: morphimF.
+ apply/eqP; rewrite (trivg_Phi (morphim_pgroup _ pP)) /= -quotientE.
+ apply/abelemP=> //; rewrite [abelian _]quotient_cents2 ?joing_subl //.
+ split=> // _ /morphimP[x Nx Px ->] /=.
+ rewrite -morphX //= coset_id // (MhoE 1 pP) joing_idr expn1.
+ by rewrite mem_gen //; apply/setUP; right; exact: mem_imset.
+rewrite -quotient_cents2 // [_ \subset 'C(_)]abP (MhoE 1 pP) gen_subG /=.
+apply/subsetP=> _ /imsetP[x Px ->]; rewrite expn1.
+have nPhi_x: x \in 'N('Phi(P)) by exact: (subsetP nPhiP).
+by rewrite coset_idr ?groupX ?morphX ?x1P ?mem_morphim.
+Qed.
+
+Lemma Phi_Mho : abelian P -> 'Phi(P) = 'Mho^1(P).
+Proof. by move=> cPP; rewrite Phi_joing (derG1P cPP) joing1G. Qed.
+
+End Frattini3.
+
+Section Frattini4.
+
+Variables (p : nat) (gT : finGroupType).
+Implicit Types (rT : finGroupType) (P G H K D : {group gT}).
+
+Lemma PhiS G H : p.-group H -> G \subset H -> 'Phi(G) \subset 'Phi(H).
+Proof.
+move=> pH sGH; rewrite (Phi_joing pH) (Phi_joing (pgroupS sGH pH)).
+by rewrite genS // setUSS ?dergS ?MhoS.
+Qed.
+
+Lemma morphim_Phi rT P D (f : {morphism D >-> rT}) :
+ p.-group P -> P \subset D -> f @* 'Phi(P) = 'Phi(f @* P).
+Proof.
+move=> pP sPD; rewrite !(@Phi_joing _ p) ?morphim_pgroup //.
+rewrite morphim_gen ?(subset_trans _ sPD) ?subUset ?der_subS ?Mho_sub //.
+by rewrite morphimU -joingE morphimR ?morphim_Mho.
+Qed.
+
+Lemma quotient_Phi P H :
+ p.-group P -> P \subset 'N(H) -> 'Phi(P) / H = 'Phi(P / H).
+Proof. exact: morphim_Phi. Qed.
+
+(* This is Aschbacher (23.2) *)
+Lemma Phi_min G H :
+ p.-group G -> G \subset 'N(H) -> p.-abelem (G / H) -> 'Phi(G) \subset H.
+Proof.
+move=> pG nHG; rewrite -trivg_Phi ?quotient_pgroup // -subG1 /=.
+by rewrite -(quotient_Phi pG) ?quotient_sub1 // (subset_trans (Phi_sub _)).
+Qed.
+
+Lemma Phi_cprod G H K :
+ p.-group G -> H \* K = G -> 'Phi(H) \* 'Phi(K) = 'Phi(G).
+Proof.
+move=> pG defG; have [_ /mulG_sub[sHG sKG] cHK] := cprodP defG.
+rewrite cprodEY /=; last by rewrite (centSS (Phi_sub _) (Phi_sub _)).
+rewrite !(Phi_joing (pgroupS _ pG)) //=.
+have /cprodP[_ <- /cent_joinEr <-] := der_cprod 1 defG.
+have /cprodP[_ <- /cent_joinEr <-] := Mho_cprod 1 defG.
+by rewrite !joingA /= -!(joingA H^`(1)) (joingC K^`(1)).
+Qed.
+
+Lemma Phi_mulg H K :
+ p.-group H -> p.-group K -> K \subset 'C(H) ->
+ 'Phi(H * K) = 'Phi(H) * 'Phi(K).
+Proof.
+move=> pH pK cHK; have defHK := cprodEY cHK.
+have [|_ -> _] := cprodP (Phi_cprod _ defHK); rewrite /= cent_joinEr //.
+by apply: pnat_dvd (dvdn_cardMg H K) _; rewrite pnat_mul; exact/andP.
+Qed.
+
+Lemma charsimpleP G :
+ reflect (G :!=: 1 /\ forall K, K :!=: 1 -> K \char G -> K :=: G)
+ (charsimple G).
+Proof.
+apply: (iffP mingroupP); rewrite char_refl andbT => [[ntG simG]].
+ by split=> // K ntK chK; apply: simG; rewrite ?ntK // char_sub.
+split=> // K /andP[ntK chK] _; exact: simG.
+Qed.
+
+End Frattini4.
+
+Section Fitting.
+
+Variable gT : finGroupType.
+Implicit Types (p : nat) (G H : {group gT}).
+
+Lemma Fitting_normal G : 'F(G) <| G.
+Proof.
+rewrite -['F(G)](bigdprodWY (erefl 'F(G))).
+elim/big_rec: _ => [|p H _ nsHG]; first by rewrite gen0 normal1.
+by rewrite -[<<_>>]joing_idr normalY ?pcore_normal.
+Qed.
+
+Lemma Fitting_sub G : 'F(G) \subset G.
+Proof. by rewrite normal_sub ?Fitting_normal. Qed.
+
+Lemma Fitting_nil G : nilpotent 'F(G).
+Proof.
+apply: (bigdprod_nil (erefl 'F(G))) => p _.
+exact: pgroup_nil (pcore_pgroup p G).
+Qed.
+
+Lemma Fitting_max G H : H <| G -> nilpotent H -> H \subset 'F(G).
+Proof.
+move=> nsHG nilH; rewrite -(Sylow_gen H) gen_subG.
+apply/bigcupsP=> P /SylowP[p _ SylP].
+case Gp: (p \in \pi(G)); last first.
+ rewrite card1_trivg ?sub1G // (card_Hall SylP).
+ rewrite part_p'nat // (pnat_dvd (cardSg (normal_sub nsHG))) //.
+ by rewrite /pnat cardG_gt0 all_predC has_pred1 Gp.
+move/nilpotent_Hall_pcore: SylP => ->{P} //.
+rewrite -(bigdprodWY (erefl 'F(G))) sub_gen //.
+rewrite -(filter_pi_of (ltnSn _)) big_filter big_mkord.
+have le_pG: p < #|G|.+1.
+ by rewrite ltnS dvdn_leq //; move: Gp; rewrite mem_primes => /and3P[].
+apply: (bigcup_max (Ordinal le_pG)) => //=.
+apply: pcore_max (pcore_pgroup _ _) _.
+exact: char_normal_trans (pcore_char p H) nsHG.
+Qed.
+
+Lemma pcore_Fitting pi G : 'O_pi('F(G)) \subset 'O_pi(G).
+Proof.
+rewrite pcore_max ?pcore_pgroup //.
+exact: char_normal_trans (pcore_char _ _) (Fitting_normal _).
+Qed.
+
+Lemma p_core_Fitting p G : 'O_p('F(G)) = 'O_p(G).
+Proof.
+apply/eqP; rewrite eqEsubset pcore_Fitting pcore_max ?pcore_pgroup //.
+apply: normalS (normal_sub (Fitting_normal _)) (pcore_normal _ _).
+exact: Fitting_max (pcore_normal _ _) (pgroup_nil (pcore_pgroup _ _)).
+Qed.
+
+Lemma nilpotent_Fitting G : nilpotent G -> 'F(G) = G.
+Proof.
+by move=> nilG; apply/eqP; rewrite eqEsubset Fitting_sub Fitting_max.
+Qed.
+
+Lemma Fitting_eq_pcore p G : 'O_p^'(G) = 1 -> 'F(G) = 'O_p(G).
+Proof.
+move=> p'G1; have /dprodP[_ /= <- _ _] := nilpotent_pcoreC p (Fitting_nil G).
+by rewrite p_core_Fitting ['O_p^'(_)](trivgP _) ?mulg1 // -p'G1 pcore_Fitting.
+Qed.
+
+Lemma FittingEgen G :
+ 'F(G) = <<\bigcup_(p < #|G|.+1 | (p : nat) \in \pi(G)) 'O_p(G)>>.
+Proof.
+apply/eqP; rewrite eqEsubset gen_subG /=.
+rewrite -{1}(bigdprodWY (erefl 'F(G))) (big_nth 0) big_mkord genS.
+ by apply/bigcupsP=> p _; rewrite -p_core_Fitting pcore_sub.
+apply/bigcupsP=> [[i /= lti]] _; set p := nth _ _ i.
+have pi_p: p \in \pi(G) by rewrite mem_nth.
+have p_dv_G: p %| #|G| by rewrite mem_primes in pi_p; case/and3P: pi_p.
+have lepG: p < #|G|.+1 by rewrite ltnS dvdn_leq.
+by rewrite (bigcup_max (Ordinal lepG)).
+Qed.
+
+End Fitting.
+
+Section FittingFun.
+
+Implicit Types gT rT : finGroupType.
+
+Lemma morphim_Fitting : GFunctor.pcontinuous Fitting.
+Proof.
+move=> gT rT G D f; apply: Fitting_max.
+ by rewrite morphim_normal ?Fitting_normal.
+by rewrite morphim_nil ?Fitting_nil.
+Qed.
+
+Lemma FittingS gT (G H : {group gT}) : H \subset G -> H :&: 'F(G) \subset 'F(H).
+Proof.
+move=> sHG; rewrite -{2}(setIidPl sHG).
+do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; exact: morphim_Fitting.
+Qed.
+
+Lemma FittingJ gT (G : {group gT}) x : 'F(G :^ x) = 'F(G) :^ x.
+Proof.
+rewrite !FittingEgen -genJ /= cardJg; symmetry; congr <<_>>.
+rewrite (big_morph (conjugate^~ x) (fun A B => conjUg A B x) (imset0 _)).
+by apply: eq_bigr => p _; rewrite pcoreJ.
+Qed.
+
+End FittingFun.
+
+Canonical Fitting_igFun := [igFun by Fitting_sub & morphim_Fitting].
+Canonical Fitting_gFun := [gFun by morphim_Fitting].
+Canonical Fitting_pgFun := [pgFun by morphim_Fitting].
+
+Section IsoFitting.
+
+Variables (gT rT : finGroupType) (G D : {group gT}) (f : {morphism D >-> rT}).
+
+Lemma Fitting_char : 'F(G) \char G. Proof. exact: gFchar. Qed.
+
+Lemma injm_Fitting : 'injm f -> G \subset D -> f @* 'F(G) = 'F(f @* G).
+Proof. exact: injmF. Qed.
+
+Lemma isog_Fitting (H : {group rT}) : G \isog H -> 'F(G) \isog 'F(H).
+Proof. exact: gFisog. Qed.
+
+End IsoFitting.
+
+Section CharSimple.
+
+Variable gT : finGroupType.
+Implicit Types (rT : finGroupType) (G H K L : {group gT}) (p : nat).
+
+Lemma minnormal_charsimple G H : minnormal H G -> charsimple H.
+Proof.
+case/mingroupP=> /andP[ntH nHG] minH.
+apply/charsimpleP; split=> // K ntK chK.
+by apply: minH; rewrite ?ntK (char_sub chK, char_norm_trans chK).
+Qed.
+
+Lemma maxnormal_charsimple G H L :
+ G <| L -> maxnormal H G L -> charsimple (G / H).
+Proof.
+case/andP=> sGL nGL /maxgroupP[/andP[/andP[sHG not_sGH] nHL] maxH].
+have nHG: G \subset 'N(H) := subset_trans sGL nHL.
+apply/charsimpleP; rewrite -subG1 quotient_sub1 //; split=> // HK ntHK chHK.
+case/(inv_quotientN _): (char_normal chHK) => [|K defHK sHK]; first exact/andP.
+case/andP; rewrite subEproper defHK => /predU1P[-> // | ltKG] nKG.
+have nHK: H <| K by rewrite /normal sHK (subset_trans (proper_sub ltKG)).
+case/negP: ntHK; rewrite defHK -subG1 quotient_sub1 ?normal_norm //.
+rewrite (maxH K) // ltKG -(quotientGK nHK) -defHK norm_quotient_pre //.
+by rewrite (char_norm_trans chHK) ?quotient_norms.
+Qed.
+
+Lemma abelem_split_dprod rT p (A B : {group rT}) :
+ p.-abelem A -> B \subset A -> exists C : {group rT}, B \x C = A.
+Proof.
+move=> abelA sBA; have [_ cAA _]:= and3P abelA.
+case/splitsP: (abelem_splits abelA sBA) => C /complP[tiBC defA].
+by exists C; rewrite dprodE // (centSS _ sBA cAA) // -defA mulG_subr.
+Qed.
+
+Lemma p_abelem_split1 rT p (A : {group rT}) x :
+ p.-abelem A -> x \in A ->
+ exists B : {group rT}, [/\ B \subset A, #|B| = #|A| %/ #[x] & <[x]> \x B = A].
+Proof.
+move=> abelA Ax; have sxA: <[x]> \subset A by rewrite cycle_subG.
+have [B defA] := abelem_split_dprod abelA sxA.
+have [_ defxB _ ti_xB] := dprodP defA.
+have sBA: B \subset A by rewrite -defxB mulG_subr.
+by exists B; split; rewrite // -defxB (TI_cardMg ti_xB) mulKn ?order_gt0.
+Qed.
+
+Lemma abelem_charsimple p G : p.-abelem G -> G :!=: 1 -> charsimple G.
+Proof.
+move=> abelG ntG; apply/charsimpleP; split=> // K ntK /charP[sKG chK].
+case/eqVproper: sKG => // /properP[sKG [x Gx notKx]].
+have ox := abelem_order_p abelG Gx (group1_contra notKx).
+have [A [sAG oA defA]] := p_abelem_split1 abelG Gx.
+case/trivgPn: ntK => y Ky nty; have Gy := subsetP sKG y Ky.
+have{nty} oy := abelem_order_p abelG Gy nty.
+have [B [sBG oB defB]] := p_abelem_split1 abelG Gy.
+have: isog A B; last case/isogP=> fAB injAB defAB.
+ rewrite (isog_abelem_card _ (abelemS sAG abelG)) (abelemS sBG) //=.
+ by rewrite oA oB ox oy.
+have: isog <[x]> <[y]>; last case/isogP=> fxy injxy /= defxy.
+ by rewrite isog_cyclic_card ?cycle_cyclic // [#|_|]oy -ox eqxx.
+have cfxA: fAB @* A \subset 'C(fxy @* <[x]>).
+ by rewrite defAB defxy; case/dprodP: defB.
+have injf: 'injm (dprodm defA cfxA).
+ by rewrite injm_dprodm injAB injxy defAB defxy; apply/eqP; case/dprodP: defB.
+case/negP: notKx; rewrite -cycle_subG -(injmSK injf) ?cycle_subG //=.
+rewrite morphim_dprodml // defxy cycle_subG /= chK //.
+have [_ {4}<- _ _] := dprodP defB; have [_ {3}<- _ _] := dprodP defA.
+by rewrite morphim_dprodm // defAB defxy.
+Qed.
+
+Lemma charsimple_dprod G : charsimple G ->
+ exists H : {group gT}, [/\ H \subset G, simple H
+ & exists2 I : {set {perm gT}}, I \subset Aut G
+ & \big[dprod/1]_(f in I) f @: H = G].
+Proof.
+case/charsimpleP=> ntG simG.
+have [H minH sHG]: {H : {group gT} | minnormal H G & H \subset G}.
+ by apply: mingroup_exists; rewrite ntG normG.
+case/mingroupP: minH => /andP[ntH nHG] minH.
+pose Iok (I : {set {perm gT}}) :=
+ (I \subset Aut G) &&
+ [exists (M : {group gT} | M <| G), \big[dprod/1]_(f in I) f @: H == M].
+have defH: (1 : {perm gT}) @: H = H.
+ apply/eqP; rewrite eqEcard card_imset ?leqnn; last exact: perm_inj.
+ by rewrite andbT; apply/subsetP=> _ /imsetP[x Hx ->]; rewrite perm1.
+have [|I] := @maxset_exists _ Iok 1.
+ rewrite /Iok sub1G; apply/existsP; exists H.
+ by rewrite /normal sHG nHG (big_pred1 1) => [|f]; rewrite ?defH /= ?inE.
+case/maxsetP=> /andP[Aut_I /exists_eq_inP[M /andP[sMG nMG] defM]] maxI.
+rewrite sub1set=> ntI; case/eqVproper: sMG => [defG | /andP[sMG not_sGM]].
+ exists H; split=> //; last by exists I; rewrite ?defM.
+ apply/mingroupP; rewrite ntH normG; split=> // N /andP[ntN nNH] sNH.
+ apply: minH => //; rewrite ntN /= -defG.
+ move: defM; rewrite (bigD1 1) //= defH; case/dprodP=> [[_ K _ ->] <- cHK _].
+ by rewrite mul_subG // cents_norm // (subset_trans cHK) ?centS.
+have defG: <<\bigcup_(f in Aut G) f @: H>> = G.
+ have sXG: \bigcup_(f in Aut G) f @: H \subset G.
+ by apply/bigcupsP=> f Af; rewrite -(im_autm Af) morphimEdom imsetS.
+ apply: simG.
+ apply: contra ntH; rewrite -!subG1; apply: subset_trans.
+ by rewrite sub_gen // (bigcup_max 1) ?group1 ?defH.
+ rewrite /characteristic gen_subG sXG; apply/forall_inP=> f Af.
+ rewrite -(autmE Af) -morphimEsub ?gen_subG ?morphim_gen // genS //.
+ rewrite morphimEsub //= autmE.
+ apply/subsetP=> _ /imsetP[_ /bigcupP[g Ag /imsetP[x Hx ->]] ->].
+ apply/bigcupP; exists (g * f); first exact: groupM.
+ by apply/imsetP; exists x; rewrite // permM.
+have [f Af sfHM]: exists2 f, f \in Aut G & ~~ (f @: H \subset M).
+ move: not_sGM; rewrite -{1}defG gen_subG; case/subsetPn=> x.
+ by case/bigcupP=> f Af fHx Mx; exists f => //; apply/subsetPn; exists x.
+case If: (f \in I).
+ by case/negP: sfHM; rewrite -(bigdprodWY defM) sub_gen // (bigcup_max f).
+case/idP: (If); rewrite -(maxI ([set f] :|: I)) ?subsetUr ?inE ?eqxx //.
+rewrite {maxI}/Iok subUset sub1set Af {}Aut_I; apply/existsP.
+have sfHG: autm Af @* H \subset G by rewrite -{4}(im_autm Af) morphimS.
+have{minH nHG} /mingroupP[/andP[ntfH nfHG] minfH]: minnormal (autm Af @* H) G.
+ apply/mingroupP; rewrite andbC -{1}(im_autm Af) morphim_norms //=.
+ rewrite -subG1 sub_morphim_pre // -kerE ker_autm subG1.
+ split=> // N /andP[ntN nNG] sNfH.
+ have sNG: N \subset G := subset_trans sNfH sfHG.
+ apply/eqP; rewrite eqEsubset sNfH sub_morphim_pre //=.
+ rewrite -(morphim_invmE (injm_autm Af)) [_ @* N]minH //=.
+ rewrite -subG1 sub_morphim_pre /= ?im_autm // morphpre_invm morphim1 subG1.
+ by rewrite ntN -{1}(im_invm (injm_autm Af)) /= {2}im_autm morphim_norms.
+ by rewrite sub_morphim_pre /= ?im_autm // morphpre_invm.
+have{minfH sfHM} tifHM: autm Af @* H :&: M = 1.
+ apply/eqP/idPn=> ntMfH; case/setIidPl: sfHM.
+ rewrite -(autmE Af) -morphimEsub //.
+ by apply: minfH; rewrite ?subsetIl // ntMfH normsI.
+have cfHM: M \subset 'C(autm Af @* H).
+ rewrite centsC (sameP commG1P trivgP) -tifHM subsetI commg_subl commg_subr.
+ by rewrite (subset_trans sMG) // (subset_trans sfHG).
+exists (autm Af @* H <*> M)%G; rewrite /normal /= join_subG sMG sfHG normsY //=.
+rewrite (bigD1 f) ?inE ?eqxx // (eq_bigl (mem I)) /= => [|g]; last first.
+ by rewrite /= !inE andbC; case: eqP => // ->.
+by rewrite defM -(autmE Af) -morphimEsub // dprodE // cent_joinEr ?eqxx.
+Qed.
+
+Lemma simple_sol_prime G : solvable G -> simple G -> prime #|G|.
+Proof.
+move=> solG /simpleP[ntG simG].
+have{solG} cGG: abelian G.
+ apply/commG1P; case/simG: (der_normal 1 G) => // /eqP/idPn[].
+ by rewrite proper_neq // (sol_der1_proper solG).
+case: (trivgVpdiv G) ntG => [-> | [p p_pr]]; first by rewrite eqxx.
+case/Cauchy=> // x Gx oxp _; move: p_pr; rewrite -oxp orderE.
+have: <[x]> <| G by rewrite -sub_abelian_normal ?cycle_subG.
+by case/simG=> -> //; rewrite cards1.
+Qed.
+
+Lemma charsimple_solvable G : charsimple G -> solvable G -> is_abelem G.
+Proof.
+case/charsimple_dprod=> H [sHG simH [I Aut_I defG]] solG.
+have p_pr: prime #|H| by exact: simple_sol_prime (solvableS sHG solG) simH.
+set p := #|H| in p_pr; apply/is_abelemP; exists p => //.
+elim/big_rec: _ (G) defG => [_ <-|f B If IH_B M defM]; first exact: abelem1.
+have [Af [[_ K _ defB] _ _ _]] := (subsetP Aut_I f If, dprodP defM).
+rewrite (dprod_abelem p defM) defB IH_B // andbT -(autmE Af) -morphimEsub //=.
+rewrite morphim_abelem ?abelemE // exponent_dvdn.
+by rewrite cyclic_abelian ?prime_cyclic.
+Qed.
+
+Lemma minnormal_solvable L G H :
+ minnormal H L -> H \subset G -> solvable G ->
+ [/\ L \subset 'N(H), H :!=: 1 & is_abelem H].
+Proof.
+move=> minH sHG solG; have /andP[ntH nHL] := mingroupp minH.
+split=> //; apply: (charsimple_solvable (minnormal_charsimple minH)).
+exact: solvableS solG.
+Qed.
+
+Lemma solvable_norm_abelem L G :
+ solvable G -> G <| L -> G :!=: 1 ->
+ exists H : {group gT}, [/\ H \subset G, H <| L, H :!=: 1 & is_abelem H].
+Proof.
+move=> solG /andP[sGL nGL] ntG.
+have [H minH sHG]: {H : {group gT} | minnormal H L & H \subset G}.
+ by apply: mingroup_exists; rewrite ntG.
+have [nHL ntH abH] := minnormal_solvable minH sHG solG.
+by exists H; split; rewrite // /normal (subset_trans sHG).
+Qed.
+
+Lemma trivg_Fitting G : solvable G -> ('F(G) == 1) = (G :==: 1).
+Proof.
+move=> solG; apply/idP/idP=> [F1|]; last first.
+ by rewrite -!subG1; apply: subset_trans; exact: Fitting_sub.
+apply/idPn=> /(solvable_norm_abelem solG (normal_refl _))[M [_ nsMG ntM]].
+case/is_abelemP=> p _ /and3P[pM _ _]; case/negP: ntM.
+by rewrite -subG1 -(eqP F1) Fitting_max ?(pgroup_nil pM).
+Qed.
+
+Lemma Fitting_pcore pi G : 'F('O_pi(G)) = 'O_pi('F(G)).
+Proof.
+apply/eqP; rewrite eqEsubset.
+rewrite (subset_trans _ (pcoreS _ (Fitting_sub _))); last first.
+ rewrite subsetI Fitting_sub Fitting_max ?Fitting_nil //.
+ by rewrite (char_normal_trans (Fitting_char _)) ?pcore_normal.
+rewrite (subset_trans _ (FittingS (pcore_sub _ _))) // subsetI pcore_sub.
+rewrite pcore_max ?pcore_pgroup //.
+by rewrite (char_normal_trans (pcore_char _ _)) ?Fitting_normal.
+Qed.
+
+End CharSimple.
+
+Section SolvablePrimeFactor.
+
+Variables (gT : finGroupType) (G : {group gT}).
+
+Lemma index_maxnormal_sol_prime (H : {group gT}) :
+ solvable G -> maxnormal H G G -> prime #|G : H|.
+Proof.
+move=> solG maxH; have nsHG := maxnormal_normal maxH.
+rewrite -card_quotient ?normal_norm // simple_sol_prime ?quotient_sol //.
+by rewrite quotient_simple.
+Qed.
+
+Lemma sol_prime_factor_exists :
+ solvable G -> G :!=: 1 -> {H : {group gT} | H <| G & prime #|G : H| }.
+Proof.
+move=> solG /ex_maxnormal_ntrivg[H maxH].
+by exists H; [exact: maxnormal_normal | exact: index_maxnormal_sol_prime].
+Qed.
+
+End SolvablePrimeFactor.
+
+Section Special.
+
+Variables (gT : finGroupType) (p : nat) (A G : {group gT}).
+
+(* This is Aschbacher (23.7) *)
+Lemma center_special_abelem : p.-group G -> special G -> p.-abelem 'Z(G).
+Proof.
+move=> pG [defPhi defG'].
+have [-> | ntG] := eqsVneq G 1; first by rewrite center1 abelem1.
+have [p_pr _ _] := pgroup_pdiv pG ntG.
+have fM: {in 'Z(G) &, {morph expgn^~ p : x y / x * y}}.
+ by move=> x y /setIP[_ /centP cxG] /setIP[/cxG cxy _]; exact: expgMn.
+rewrite abelemE //= center_abelian; apply/exponentP=> /= z Zz.
+apply: (@kerP _ _ _ (Morphism fM)) => //; apply: subsetP z Zz.
+rewrite -{1}defG' gen_subG; apply/subsetP=> _ /imset2P[x y Gx Gy ->].
+have Zxy: [~ x, y] \in 'Z(G) by rewrite -defG' mem_commg.
+have Zxp: x ^+ p \in 'Z(G).
+ rewrite -defPhi (Phi_joing pG) (MhoE 1 pG) joing_idr mem_gen // !inE.
+ by rewrite expn1 orbC (mem_imset (expgn^~ p)).
+rewrite mem_morphpre /= ?defG' ?Zxy // inE -commXg; last first.
+ by red; case/setIP: Zxy => _ /centP->.
+by apply/commgP; red; case/setIP: Zxp => _ /centP->.
+Qed.
+
+Lemma exponent_special : p.-group G -> special G -> exponent G %| p ^ 2.
+Proof.
+move=> pG spG; have [defPhi _] := spG.
+have /and3P[_ _ expZ] := center_special_abelem pG spG.
+apply/exponentP=> x Gx; rewrite expgM (exponentP expZ) // -defPhi.
+by rewrite (Phi_joing pG) mem_gen // inE orbC (Mho_p_elt 1) ?(mem_p_elt pG).
+Qed.
+
+(* Aschbacher 24.7 (replaces Gorenstein 5.3.7) *)
+Theorem abelian_charsimple_special :
+ p.-group G -> coprime #|G| #|A| -> [~: G, A] = G ->
+ \bigcup_(H : {group gT} | (H \char G) && abelian H) H \subset 'C(A) ->
+ special G /\ 'C_G(A) = 'Z(G).
+Proof.
+move=> pG coGA defG /bigcupsP cChaA.
+have cZA: 'Z(G) \subset 'C_G(A).
+ by rewrite subsetI center_sub cChaA // center_char center_abelian.
+have cChaG (H : {group gT}): H \char G -> abelian H -> H \subset 'Z(G).
+ move=> chH abH; rewrite subsetI char_sub //= centsC -defG.
+ rewrite comm_norm_cent_cent ?(char_norm chH) -?commg_subl ?defG //.
+ by rewrite centsC cChaA ?chH.
+have cZ2GG: [~: 'Z_2(G), G, G] = 1.
+ by apply/commG1P; rewrite (subset_trans (ucn_comm 1 G)) // ucn1 subsetIr.
+have{cZ2GG} cG'Z: 'Z_2(G) \subset 'C(G^`(1)).
+ by rewrite centsC; apply/commG1P; rewrite three_subgroup // (commGC G).
+have{cG'Z} sZ2G'_Z: 'Z_2(G) :&: G^`(1) \subset 'Z(G).
+ apply: cChaG; first by rewrite charI ?ucn_char ?der_char.
+ by rewrite /abelian subIset // (subset_trans cG'Z) // centS ?subsetIr.
+have{sZ2G'_Z} sG'Z: G^`(1) \subset 'Z(G).
+ rewrite der1_min ?gFnorm //; apply/derG1P.
+ have /TI_center_nil: nilpotent (G / 'Z(G)) := quotient_nil _ (pgroup_nil pG).
+ apply; first exact: gFnormal; rewrite /= setIC -ucn1 -ucn_central.
+ rewrite -quotient_der ?gFnorm // -quotientGI ?ucn_subS ?quotientS1 //=.
+ by rewrite ucn1.
+have sCG': 'C_G(A) \subset G^`(1).
+ rewrite -quotient_sub1 //; last by rewrite subIset // char_norm ?der_char.
+ rewrite (subset_trans (quotient_subcent _ G A)) //= -[G in G / _]defG.
+ have nGA: A \subset 'N(G) by rewrite -commg_subl defG.
+ rewrite quotientR ?(char_norm_trans (der_char _ _)) ?normG //.
+ rewrite coprime_abel_cent_TI ?quotient_norms ?coprime_morph //.
+ exact: sub_der1_abelian.
+have defZ: 'Z(G) = G^`(1) by apply/eqP; rewrite eqEsubset (subset_trans cZA).
+split; last by apply/eqP; rewrite eqEsubset cZA defZ sCG'.
+split=> //; apply/eqP; rewrite eqEsubset defZ (Phi_joing pG) joing_subl.
+have:= pG; rewrite -pnat_exponent => /p_natP[n expGpn].
+rewrite join_subG subxx andbT /= -defZ -(subnn n.-1).
+elim: {2}n.-1 => [|m IHm].
+ rewrite (MhoE _ pG) gen_subG; apply/subsetP=> _ /imsetP[x Gx ->].
+ rewrite subn0 -subn1 -add1n -maxnE maxnC maxnE expnD.
+ by rewrite expgM -expGpn expg_exponent ?groupX ?group1.
+rewrite cChaG ?Mho_char //= (MhoE _ pG) /abelian cent_gen gen_subG.
+apply/centsP=> _ /imsetP[x Gx ->] _ /imsetP[y Gy ->].
+move: sG'Z; rewrite subsetI centsC => /andP[_ /centsP cGG'].
+apply/commgP; rewrite {1}expnSr expgM.
+rewrite commXg -?commgX; try by apply: cGG'; rewrite ?mem_commg ?groupX.
+apply/commgP; rewrite subsetI Mho_sub centsC in IHm.
+apply: (centsP IHm); first by rewrite groupX.
+rewrite -add1n -(addn1 m) subnDA -maxnE maxnC maxnE.
+rewrite -expgM -expnSr -addSn expnD expgM groupX //=.
+by rewrite Mho_p_elt ?(mem_p_elt pG).
+Qed.
+
+End Special.
+
+Section Extraspecial.
+
+Variables (p : nat) (gT rT : finGroupType).
+Implicit Types D E F G H K M R S T U : {group gT}.
+
+Section Basic.
+
+Variable S : {group gT}.
+Hypotheses (pS : p.-group S) (esS : extraspecial S).
+
+Let pZ : p.-group 'Z(S) := pgroupS (center_sub S) pS.
+Lemma extraspecial_prime : prime p.
+Proof.
+by case: esS => _ /prime_gt1; rewrite cardG_gt1; case/(pgroup_pdiv pZ).
+Qed.
+
+Lemma card_center_extraspecial : #|'Z(S)| = p.
+Proof. by apply/eqP; apply: (pgroupP pZ); case: esS. Qed.
+
+Lemma min_card_extraspecial : #|S| >= p ^ 3.
+Proof.
+have p_gt1 := prime_gt1 extraspecial_prime.
+rewrite leqNgt (card_pgroup pS) ltn_exp2l // ltnS.
+case: esS => [[_ defS']]; apply: contraL => /(p2group_abelian pS)/derG1P S'1.
+by rewrite -defS' S'1 cards1.
+Qed.
+
+End Basic.
+
+Lemma card_p3group_extraspecial E :
+ prime p -> #|E| = (p ^ 3)%N -> #|'Z(E)| = p -> extraspecial E.
+Proof.
+move=> p_pr oEp3 oZp; have p_gt0 := prime_gt0 p_pr.
+have pE: p.-group E by rewrite /pgroup oEp3 pnat_exp pnat_id.
+have pEq: p.-group (E / 'Z(E))%g by rewrite quotient_pgroup.
+have /andP[sZE nZE] := center_normal E.
+have oEq: #|E / 'Z(E)|%g = (p ^ 2)%N.
+ by rewrite card_quotient -?divgS // oEp3 oZp expnS mulKn.
+have cEEq: abelian (E / 'Z(E))%g by exact: card_p2group_abelian oEq.
+have not_cEE: ~~ abelian E.
+ have: #|'Z(E)| < #|E| by rewrite oEp3 oZp (ltn_exp2l 1) ?prime_gt1.
+ by apply: contraL => cEE; rewrite -leqNgt subset_leq_card // subsetI subxx.
+have defE': E^`(1) = 'Z(E).
+ apply/eqP; rewrite eqEsubset der1_min //=; apply: contraR not_cEE => not_sE'Z.
+ apply/commG1P/(TI_center_nil (pgroup_nil pE) (der_normal 1 _)).
+ by rewrite setIC prime_TIg ?oZp.
+split; [split=> // | by rewrite oZp]; apply/eqP.
+rewrite eqEsubset andbC -{1}defE' {1}(Phi_joing pE) joing_subl.
+rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) //.
+rewrite subG1 /= (quotient_Phi pE) //= (trivg_Phi pEq); apply/abelemP=> //.
+split=> // Zx EqZx; apply/eqP; rewrite -order_dvdn /order.
+rewrite (card_pgroup (mem_p_elt pEq EqZx)) (@dvdn_exp2l _ _ 1) //.
+rewrite leqNgt -pfactor_dvdn // -oEq; apply: contra not_cEE => sEqZx.
+rewrite cyclic_center_factor_abelian //; apply/cyclicP.
+exists Zx; apply/eqP; rewrite eq_sym eqEcard cycle_subG EqZx -orderE.
+exact: dvdn_leq sEqZx.
+Qed.
+
+Lemma p3group_extraspecial G :
+ p.-group G -> ~~ abelian G -> logn p #|G| <= 3 -> extraspecial G.
+Proof.
+move=> pG not_cGG; have /andP[sZG nZG] := center_normal G.
+have ntG: G :!=: 1 by apply: contraNneq not_cGG => ->; exact: abelian1.
+have ntZ: 'Z(G) != 1 by rewrite (center_nil_eq1 (pgroup_nil pG)).
+have [p_pr _ [n oG]] := pgroup_pdiv pG ntG; rewrite oG pfactorK //.
+have [_ _ [m oZ]] := pgroup_pdiv (pgroupS sZG pG) ntZ.
+have lt_m1_n: m.+1 < n.
+ suffices: 1 < logn p #|(G / 'Z(G))|.
+ rewrite card_quotient // -divgS // logn_div ?cardSg //.
+ by rewrite oG oZ !pfactorK // ltn_subRL addn1.
+ rewrite ltnNge; apply: contra not_cGG => cycGs.
+ apply: cyclic_center_factor_abelian; rewrite (dvdn_prime_cyclic p_pr) //.
+ by rewrite (card_pgroup (quotient_pgroup _ pG)) (dvdn_exp2l _ cycGs).
+rewrite -{lt_m1_n}(subnKC lt_m1_n) !addSn !ltnS leqn0 in oG *.
+case: m => // in oZ oG * => /eqP n2; rewrite {n}n2 in oG.
+exact: card_p3group_extraspecial oZ.
+Qed.
+
+Lemma extraspecial_nonabelian G : extraspecial G -> ~~ abelian G.
+Proof.
+case=> [[_ defG'] oZ]; rewrite /abelian (sameP commG1P eqP).
+by rewrite -derg1 defG' -cardG_gt1 prime_gt1.
+Qed.
+
+Lemma exponent_2extraspecial G : 2.-group G -> extraspecial G -> exponent G = 4.
+Proof.
+move=> p2G esG; have [spG _] := esG.
+case/dvdn_pfactor: (exponent_special p2G spG) => // k.
+rewrite leq_eqVlt ltnS => /predU1P[-> // | lek1] expG.
+case/negP: (extraspecial_nonabelian esG).
+by rewrite (@abelem_abelian _ 2) ?exponent2_abelem // expG pfactor_dvdn.
+Qed.
+
+Lemma injm_special D G (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> special G -> special (f @* G).
+Proof.
+move=> injf sGD [defPhiG defG'].
+by rewrite /special -morphim_der // -injm_Phi // defPhiG defG' injm_center.
+Qed.
+
+Lemma injm_extraspecial D G (f : {morphism D >-> rT}) :
+ 'injm f -> G \subset D -> extraspecial G -> extraspecial (f @* G).
+Proof.
+move=> injf sGD [spG ZG_pr]; split; first exact: injm_special spG.
+by rewrite -injm_center // card_injm // subIset ?sGD.
+Qed.
+
+Lemma isog_special G (R : {group rT}) :
+ G \isog R -> special G -> special R.
+Proof. by case/isogP=> f injf <-; exact: injm_special. Qed.
+
+Lemma isog_extraspecial G (R : {group rT}) :
+ G \isog R -> extraspecial G -> extraspecial R.
+Proof. by case/isogP=> f injf <-; exact: injm_extraspecial. Qed.
+
+Lemma cprod_extraspecial G H K :
+ p.-group G -> H \* K = G -> H :&: K = 'Z(H) ->
+ extraspecial H -> extraspecial K -> extraspecial G.
+Proof.
+move=> pG defG ziHK [[PhiH defH'] ZH_pr] [[PhiK defK'] ZK_pr].
+have [_ defHK cHK]:= cprodP defG.
+have sZHK: 'Z(H) \subset 'Z(K).
+ by rewrite subsetI -{1}ziHK subsetIr subIset // centsC cHK.
+have{sZHK} defZH: 'Z(H) = 'Z(K).
+ by apply/eqP; rewrite eqEcard sZHK leq_eqVlt eq_sym -dvdn_prime2 ?cardSg.
+have defZ: 'Z(G) = 'Z(K).
+ by case/cprodP: (center_cprod defG) => /= _ <- _; rewrite defZH mulGid.
+split; first split; rewrite defZ //.
+ by have /cprodP[_ <- _] := Phi_cprod pG defG; rewrite PhiH PhiK defZH mulGid.
+by have /cprodP[_ <- _] := der_cprod 1 defG; rewrite defH' defK' defZH mulGid.
+Qed.
+
+(* Lemmas bundling Aschbacher (23.10) with (19.1), (19.2), (19.12) and (20.8) *)
+Section ExtraspecialFormspace.
+
+Variable G : {group gT}.
+Hypotheses (pG : p.-group G) (esG : extraspecial G).
+
+Let p_pr := extraspecial_prime pG esG.
+Let oZ := card_center_extraspecial pG esG.
+Let p_gt1 := prime_gt1 p_pr.
+Let p_gt0 := prime_gt0 p_pr.
+
+(* This encasulates Aschbacher (23.10)(1). *)
+Lemma cent1_extraspecial_maximal x :
+ x \in G -> x \notin 'Z(G) -> maximal 'C_G[x] G.
+Proof.
+move=> Gx notZx; pose f y := [~ x, y]; have [[_ defG'] prZ] := esG.
+have{defG'} fZ y: y \in G -> f y \in 'Z(G).
+ by move=> Gy; rewrite -defG' mem_commg.
+have fM: {in G &, {morph f : y z / y * z}}%g.
+ move=> y z Gy Gz; rewrite {1}/f commgMJ conjgCV -conjgM (conjg_fixP _) //.
+ rewrite (sameP commgP cent1P); apply: subsetP (fZ y Gy).
+ by rewrite subIset // orbC -cent_set1 centS // sub1set !(groupM, groupV).
+pose fm := Morphism fM.
+have fmG: fm @* G = 'Z(G).
+ have sfmG: fm @* G \subset 'Z(G).
+ by apply/subsetP=> _ /morphimP[z _ Gz ->]; exact: fZ.
+ apply/eqP; rewrite eqEsubset sfmG; apply: contraR notZx => /(prime_TIg prZ).
+ rewrite (setIidPr _) // => fmG1; rewrite inE Gx; apply/centP=> y Gy.
+ by apply/commgP; rewrite -in_set1 -[[set _]]fmG1; exact: mem_morphim.
+have ->: 'C_G[x] = 'ker fm.
+ apply/setP=> z; rewrite inE (sameP cent1P commgP) !inE.
+ by rewrite -invg_comm eq_invg_mul mulg1.
+rewrite p_index_maximal ?subsetIl // -card_quotient ?ker_norm //.
+by rewrite (card_isog (first_isog fm)) /= fmG.
+Qed.
+
+(* This is the tranposition of the hyperplane dimension theorem (Aschbacher *)
+(* (19.1)) to subgroups of an extraspecial group. *)
+Lemma subcent1_extraspecial_maximal U x :
+ U \subset G -> x \in G :\: 'C(U) -> maximal 'C_U[x] U.
+Proof.
+move=> sUG /setDP[Gx not_cUx]; apply/maxgroupP; split=> [|H ltHU sCxH].
+ by rewrite /proper subsetIl subsetI subxx sub_cent1.
+case/andP: ltHU => sHU not_sHU; have sHG := subset_trans sHU sUG.
+apply/eqP; rewrite eqEsubset sCxH subsetI sHU /= andbT.
+apply: contraR not_sHU => not_sHCx.
+have maxCx: maximal 'C_G[x] G.
+ rewrite cent1_extraspecial_maximal //; apply: contra not_cUx.
+ by rewrite inE Gx; exact: subsetP (centS sUG) _.
+have nsCx := p_maximal_normal pG maxCx.
+rewrite -(setIidPl sUG) -(mulg_normal_maximal nsCx maxCx sHG) ?subsetI ?sHG //.
+by rewrite -group_modr //= setIA (setIidPl sUG) mul_subG.
+Qed.
+
+(* This is the tranposition of the orthogonal subspace dimension theorem *)
+(* (Aschbacher (19.2)) to subgroups of an extraspecial group. *)
+Lemma card_subcent_extraspecial U :
+ U \subset G -> #|'C_G(U)| = (#|'Z(G) :&: U| * #|G : U|)%N.
+Proof.
+move=> sUG; rewrite setIAC (setIidPr sUG).
+elim: {U}_.+1 {-2}U (ltnSn #|U|) sUG => // m IHm U leUm sUG.
+have [cUG | not_cUG]:= orP (orbN (G \subset 'C(U))).
+ by rewrite !(setIidPl _) ?Lagrange // centsC.
+have{not_cUG} [x Gx not_cUx] := subsetPn not_cUG.
+pose W := 'C_U[x]; have sCW_G: 'C_G(W) \subset G := subsetIl G _.
+have maxW: maximal W U by rewrite subcent1_extraspecial_maximal // inE not_cUx.
+have nsWU: W <| U := p_maximal_normal (pgroupS sUG pG) maxW.
+have ltWU: W \proper U by exact: maxgroupp maxW.
+have [sWU [u Uu notWu]] := properP ltWU; have sWG := subset_trans sWU sUG.
+have defU: W * <[u]> = U by rewrite (mulg_normal_maximal nsWU) ?cycle_subG.
+have iCW_CU: #|'C_G(W) : 'C_G(U)| = p.
+ rewrite -defU centM cent_cycle setIA /=; rewrite inE Uu cent1C in notWu.
+ apply: p_maximal_index (pgroupS sCW_G pG) _.
+ apply: subcent1_extraspecial_maximal sCW_G _.
+ rewrite inE andbC (subsetP sUG) //= -sub_cent1.
+ by apply/subsetPn; exists x; rewrite // inE Gx -sub_cent1 subsetIr.
+apply/eqP; rewrite -(eqn_pmul2r p_gt0) -{1}iCW_CU Lagrange ?setIS ?centS //.
+rewrite IHm ?(leq_trans (proper_card ltWU)) // -setIA -mulnA.
+rewrite -(Lagrange_index sUG sWU) (p_maximal_index (pgroupS sUG pG)) //=.
+by rewrite -cent_set1 (setIidPr (centS _)) ?sub1set.
+Qed.
+
+(* This is the tranposition of the proof that a singular vector is contained *)
+(* in a hyperbolic plane (Aschbacher (19.12)) to subgroups of an extraspecial *)
+(* group. *)
+Lemma split1_extraspecial x :
+ x \in G :\: 'Z(G) ->
+ {E : {group gT} & {R : {group gT} |
+ [/\ #|E| = (p ^ 3)%N /\ #|R| = #|G| %/ p ^ 2,
+ E \* R = G /\ E :&: R = 'Z(E),
+ 'Z(E) = 'Z(G) /\ 'Z(R) = 'Z(G),
+ extraspecial E /\ x \in E
+ & if abelian R then R :=: 'Z(G) else extraspecial R]}}.
+Proof.
+case/setDP=> Gx notZx; rewrite inE Gx /= in notZx.
+have [[defPhiG defG'] prZ] := esG.
+have maxCx: maximal 'C_G[x] G.
+ by rewrite subcent1_extraspecial_maximal // inE notZx.
+pose y := repr (G :\: 'C[x]).
+have [Gy not_cxy]: y \in G /\ y \notin 'C[x].
+ move/maxgroupp: maxCx => /properP[_ [t Gt not_cyt]].
+ by apply/setDP; apply: (mem_repr t); rewrite !inE Gt andbT in not_cyt *.
+pose E := <[x]> <*> <[y]>; pose R := 'C_G(E).
+exists [group of E]; exists [group of R] => /=.
+have sEG: E \subset G by rewrite join_subG !cycle_subG Gx.
+have [Ex Ey]: x \in E /\ y \in E by rewrite !mem_gen // inE cycle_id ?orbT.
+have sZE: 'Z(G) \subset E.
+ rewrite (('Z(G) =P E^`(1)) _) ?der_sub // eqEsubset -{2}defG' dergS // andbT.
+ apply: contraR not_cxy => /= not_sZE'.
+ rewrite (sameP cent1P commgP) -in_set1 -[[set 1]](prime_TIg prZ not_sZE').
+ by rewrite /= -defG' inE !mem_commg.
+have ziER: E :&: R = 'Z(E) by rewrite setIA (setIidPl sEG).
+have cER: R \subset 'C(E) by rewrite subsetIr.
+have iCxG: #|G : 'C_G[x]| = p by exact: p_maximal_index.
+have maxR: maximal R 'C_G[x].
+ rewrite /R centY !cent_cycle setIA.
+ rewrite subcent1_extraspecial_maximal ?subsetIl // inE Gy andbT -sub_cent1.
+ by apply/subsetPn; exists x; rewrite 1?cent1C // inE Gx cent1id.
+have sRCx: R \subset 'C_G[x] by rewrite -cent_cycle setIS ?centS ?joing_subl.
+have sCxG: 'C_G[x] \subset G by rewrite subsetIl.
+have sRG: R \subset G by rewrite subsetIl.
+have iRCx: #|'C_G[x] : R| = p by rewrite (p_maximal_index (pgroupS sCxG pG)).
+have defG: E * R = G.
+ rewrite -cent_joinEr //= -/R joingC joingA.
+ have cGx_x: <[x]> \subset 'C_G[x] by rewrite cycle_subG inE Gx cent1id.
+ have nsRcx := p_maximal_normal (pgroupS sCxG pG) maxR.
+ rewrite (norm_joinEr (subset_trans cGx_x (normal_norm nsRcx))).
+ rewrite (mulg_normal_maximal nsRcx) //=; last first.
+ by rewrite centY !cent_cycle cycle_subG !in_setI Gx cent1id cent1C.
+ have nsCxG := p_maximal_normal pG maxCx.
+ have syG: <[y]> \subset G by rewrite cycle_subG.
+ rewrite (norm_joinEr (subset_trans syG (normal_norm nsCxG))).
+ by rewrite (mulg_normal_maximal nsCxG) //= cycle_subG inE Gy.
+have defZR: 'Z(R) = 'Z(G) by rewrite -['Z(R)]setIA -centM defG.
+have defZE: 'Z(E) = 'Z(G).
+ by rewrite -defG -center_prod ?mulGSid //= -ziER subsetI center_sub defZR sZE.
+have [n oG] := p_natP pG.
+have n_gt1: n > 1.
+ by rewrite ltnW // -(@leq_exp2l p) // -oG min_card_extraspecial.
+have oR: #|R| = (p ^ n.-2)%N.
+ apply/eqP; rewrite -(divg_indexS sRCx) iRCx /= -(divg_indexS sCxG) iCxG /= oG.
+ by rewrite -{1}(subnKC n_gt1) subn2 !expnS !mulKn.
+have oE: #|E| = (p ^ 3)%N.
+ apply/eqP; rewrite -(@eqn_pmul2r #|R|) ?cardG_gt0 // mul_cardG defG ziER.
+ by rewrite defZE oZ oG -{1}(subnKC n_gt1) oR -expnSr -expnD subn2.
+rewrite cprodE // oR oG -expnB ?subn2 //; split=> //.
+ by split=> //; apply: card_p3group_extraspecial _ oE _; rewrite // defZE.
+case: ifP => [cRR | not_cRR]; first by rewrite -defZR (center_idP _).
+split; rewrite /special defZR //.
+have ntR': R^`(1) != 1 by rewrite (sameP eqP commG1P) -abelianE not_cRR.
+have pR: p.-group R := pgroupS sRG pG.
+have pR': p.-group R^`(1) := pgroupS (der_sub 1 _) pR.
+have defR': R^`(1) = 'Z(G).
+ apply/eqP; rewrite eqEcard -{1}defG' dergS //= oZ.
+ by have [_ _ [k ->]]:= pgroup_pdiv pR' ntR'; rewrite (leq_exp2l 1).
+split=> //; apply/eqP; rewrite eqEsubset -{1}defPhiG -defR' (PhiS pG) //=.
+by rewrite (Phi_joing pR) joing_subl.
+Qed.
+
+(* This is the tranposition of the proof that the dimension of any maximal *)
+(* totally singular subspace is equal to the Witt index (Aschbacher (20.8)), *)
+(* to subgroups of an extraspecial group (in a slightly more general form, *)
+(* since we allow for p != 2). *)
+(* Note that Aschbacher derives this from the Witt lemma, which we avoid. *)
+Lemma pmaxElem_extraspecial : 'E*_p(G) = 'E_p^('r_p(G))(G).
+Proof.
+have sZmax: {in 'E*_p(G), forall E, 'Z(G) \subset E}.
+ move=> E maxE; have defE := pmaxElem_LdivP p_pr maxE.
+ have abelZ: p.-abelem 'Z(G) by rewrite prime_abelem ?oZ.
+ rewrite -(Ohm1_id abelZ) (OhmE 1 (abelem_pgroup abelZ)) gen_subG -defE.
+ by rewrite setSI // setIS ?centS // -defE !subIset ?subxx.
+suffices card_max: {in 'E*_p(G) &, forall E F, #|E| <= #|F| }.
+ have EprGmax: 'E_p^('r_p(G))(G) \subset 'E*_p(G) := p_rankElem_max p G.
+ have [E EprE]:= p_rank_witness p G; have maxE := subsetP EprGmax E EprE.
+ apply/eqP; rewrite eqEsubset EprGmax andbT; apply/subsetP=> F maxF.
+ rewrite inE; have [-> _]:= pmaxElemP maxF; have [_ _ <-]:= pnElemP EprE.
+ by apply/eqP; congr (logn p _); apply/eqP; rewrite eqn_leq !card_max.
+move=> E F maxE maxF; set U := E :&: F.
+have [sUE sUF]: U \subset E /\ U \subset F by apply/andP; rewrite -subsetI.
+have sZU: 'Z(G) \subset U by rewrite subsetI !sZmax.
+have [EpE _]:= pmaxElemP maxE; have{EpE} [sEG abelE] := pElemP EpE.
+have [EpF _]:= pmaxElemP maxF; have{EpF} [sFG abelF] := pElemP EpF.
+have [V] := abelem_split_dprod abelE sUE; case/dprodP=> _ defE cUV tiUV.
+have [W] := abelem_split_dprod abelF sUF; case/dprodP=> _ defF _ tiUW.
+have [sVE sWF]: V \subset E /\ W \subset F by rewrite -defE -defF !mulG_subr.
+have [sVG sWG] := (subset_trans sVE sEG, subset_trans sWF sFG).
+rewrite -defE -defF !TI_cardMg // leq_pmul2l ?cardG_gt0 //.
+rewrite -(leq_pmul2r (cardG_gt0 'C_G(W))) mul_cardG.
+rewrite card_subcent_extraspecial // mulnCA Lagrange // mulnC.
+rewrite leq_mul ?subset_leq_card //; last by rewrite mul_subG ?subsetIl.
+apply: subset_trans (sub1G _); rewrite -tiUV !subsetI subsetIl subIset ?sVE //=.
+rewrite -(pmaxElem_LdivP p_pr maxF) -defF centM -!setIA -(setICA 'C(W)).
+rewrite setIC setIA setIS // subsetI cUV sub_LdivT.
+by case/and3P: (abelemS sVE abelE).
+Qed.
+
+End ExtraspecialFormspace.
+
+(* This is B & G, Theorem 4.15, as done in Aschbacher (23.8) *)
+Lemma critical_extraspecial R S :
+ p.-group R -> S \subset R -> extraspecial S -> [~: S, R] \subset S^`(1) ->
+ S \* 'C_R(S) = R.
+Proof.
+move=> pR sSR esS sSR_S'; have [[defPhi defS'] _] := esS.
+have [pS [sPS nPS]] := (pgroupS sSR pR, andP (Phi_normal S : 'Phi(S) <| S)).
+have{esS} oZS: #|'Z(S)| = p := card_center_extraspecial pS esS.
+have nSR: R \subset 'N(S) by rewrite -commg_subl (subset_trans sSR_S') ?der_sub.
+have nsCR: 'C_R(S) <| R by rewrite (normalGI nSR) ?cent_normal.
+have nCS: S \subset 'N('C_R(S)) by rewrite cents_norm // centsC subsetIr.
+rewrite cprodE ?subsetIr //= -{2}(quotientGK nsCR) normC -?quotientK //.
+congr (_ @*^-1 _); apply/eqP; rewrite eqEcard quotientS //=.
+rewrite -(card_isog (second_isog nCS)) setIAC (setIidPr sSR) /= -/'Z(S) -defPhi.
+rewrite -ker_conj_aut (card_isog (first_isog_loc _ nSR)) //=; set A := _ @* R.
+have{pS} abelSb := Phi_quotient_abelem pS; have [pSb cSSb _] := and3P abelSb.
+have [/= Xb defSb oXb] := grank_witness (S / 'Phi(S)).
+pose X := (repr \o val : coset_of _ -> gT) @: Xb.
+have sXS: X \subset S; last have nPX := subset_trans sXS nPS.
+ apply/subsetP=> x; case/imsetP=> xb Xxb ->; have nPx := repr_coset_norm xb.
+ rewrite -sub1set -(quotientSGK _ sPS) ?sub1set ?quotient_set1 //= sub1set.
+ by rewrite coset_reprK -defSb mem_gen.
+have defS: <<X>> = S.
+ apply: Phi_nongen; apply/eqP; rewrite eqEsubset join_subG sPS sXS -joing_idr.
+ rewrite -genM_join sub_gen // -quotientSK ?quotient_gen // -defSb genS //.
+ apply/subsetP=> xb Xxb; apply/imsetP; rewrite (setIidPr nPX).
+ by exists (repr xb); rewrite /= ?coset_reprK //; exact: mem_imset.
+pose f (a : {perm gT}) := [ffun x => if x \in X then x^-1 * a x else 1].
+have injf: {in A &, injective f}.
+ move=> _ _ /morphimP[y nSy Ry ->] /morphimP[z nSz Rz ->].
+ move/ffunP=> eq_fyz; apply: (@eq_Aut _ S); rewrite ?Aut_aut //= => x Sx.
+ rewrite !norm_conj_autE //; apply: canRL (conjgKV z) _; rewrite -conjgM.
+ rewrite /conjg -(centP _ x Sx) ?mulKg {x Sx}// -defS cent_gen -sub_cent1.
+ apply/subsetP=> x Xx; have Sx := subsetP sXS x Xx.
+ move/(_ x): eq_fyz; rewrite !ffunE Xx !norm_conj_autE // => /mulgI xy_xz.
+ by rewrite cent1C inE conjg_set1 conjgM xy_xz conjgK.
+have sfA_XS': f @: A \subset pffun_on 1 X S^`(1).
+ apply/subsetP=> _ /imsetP[_ /morphimP[y nSy Ry ->] ->].
+ apply/pffun_onP; split=> [|_ /imageP[x /= Xx ->]].
+ by apply/subsetP=> x; apply: contraR; rewrite ffunE => /negPf->.
+ have Sx := subsetP sXS x Xx.
+ by rewrite ffunE Xx norm_conj_autE // (subsetP sSR_S') ?mem_commg.
+rewrite -(card_in_imset injf) (leq_trans (subset_leq_card sfA_XS')) // defS'.
+rewrite card_pffun_on (card_pgroup pSb) -rank_abelem -?grank_abelian // -oXb.
+by rewrite -oZS ?leq_pexp2l ?cardG_gt0 ?leq_imset_card.
+Qed.
+
+(* This is part of Aschbacher (23.13) and (23.14). *)
+Theorem extraspecial_structure S : p.-group S -> extraspecial S ->
+ {Es | all (fun E => (#|E| == p ^ 3)%N && ('Z(E) == 'Z(S))) Es
+ & \big[cprod/1%g]_(E <- Es) E \* 'Z(S) = S}.
+Proof.
+elim: {S}_.+1 {-2}S (ltnSn #|S|) => // m IHm S leSm pS esS.
+have [x Z'x]: {x | x \in S :\: 'Z(S)}.
+ apply/sigW/set0Pn; rewrite -subset0 subDset setU0.
+ apply: contra (extraspecial_nonabelian esS) => sSZ.
+ exact: abelianS sSZ (center_abelian S).
+have [E [R [[oE oR]]]]:= split1_extraspecial pS esS Z'x.
+case=> defS _ [defZE defZR] _; case: ifP => [_ defR | _ esR].
+ by exists [:: E]; rewrite /= ?oE ?defZE ?eqxx // big_seq1 -defR.
+have sRS: R \subset S by case/cprodP: defS => _ <- _; rewrite mulG_subr.
+have [|Es esEs defR] := IHm _ _ (pgroupS sRS pS) esR.
+ rewrite oR (leq_trans (ltn_Pdiv _ _)) ?cardG_gt0 // (ltn_exp2l 0) //.
+ exact: prime_gt1 (extraspecial_prime pS esS).
+exists (E :: Es); first by rewrite /= oE defZE !eqxx -defZR.
+by rewrite -defZR big_cons -cprodA defR.
+Qed.
+
+Section StructureCorollaries.
+
+Variable S : {group gT}.
+Hypotheses (pS : p.-group S) (esS : extraspecial S).
+
+Let p_pr := extraspecial_prime pS esS.
+Let oZ := card_center_extraspecial pS esS.
+
+(* This is Aschbacher (23.10)(2). *)
+Lemma card_extraspecial : {n | n > 0 & #|S| = (p ^ n.*2.+1)%N}.
+Proof.
+exists (logn p #|S|)./2.
+ rewrite half_gt0 ltnW // -(leq_exp2l _ _ (prime_gt1 p_pr)) -card_pgroup //.
+ exact: min_card_extraspecial.
+have [Es] := extraspecial_structure pS esS.
+elim: Es {3 4 5}S => [_ _ <-| E s IHs T] /=.
+ by rewrite big_nil cprod1g oZ (pfactorK 1).
+rewrite -andbA big_cons -cprodA; case/and3P; move/eqP=> oEp3; move/eqP=> defZE.
+move/IHs=> {IHs}IHs; case/cprodP=> [[_ U _ defU]]; rewrite defU => defT cEU.
+rewrite -(mulnK #|T| (cardG_gt0 (E :&: U))) -defT -mul_cardG /=.
+have ->: E :&: U = 'Z(S).
+ apply/eqP; rewrite eqEsubset subsetI -{1 2}defZE subsetIl setIS //=.
+ by case/cprodP: defU => [[V _ -> _]] <- _; exact: mulG_subr.
+rewrite (IHs U) // oEp3 oZ -expnD addSn expnS mulKn ?prime_gt0 //.
+by rewrite pfactorK //= uphalf_double.
+Qed.
+
+Lemma Aut_extraspecial_full : Aut_in (Aut S) 'Z(S) \isog Aut 'Z(S).
+Proof.
+have [p_gt1 p_gt0] := (prime_gt1 p_pr, prime_gt0 p_pr).
+have [Es] := extraspecial_structure pS esS.
+elim: Es S oZ => [T _ _ <-| E s IHs T oZT] /=.
+ rewrite big_nil cprod1g (center_idP (center_abelian T)).
+ by apply/Aut_sub_fullP=> // g injg gZ; exists g.
+rewrite -andbA big_cons -cprodA; case/and3P; move/eqP=> oE; move/eqP=> defZE.
+move=> es_s; case/cprodP=> [[_ U _ defU]]; rewrite defU => defT cEU.
+have sUT: U \subset T by rewrite -defT mulG_subr.
+have sZU: 'Z(T) \subset U.
+ by case/cprodP: defU => [[V _ -> _] <- _]; exact: mulG_subr.
+have defZU: 'Z(E) = 'Z(U).
+ apply/eqP; rewrite eqEsubset defZE subsetI sZU subIset ?centS ?orbT //=.
+ by rewrite subsetI subIset ?sUT //= -defT centM setSI.
+apply: (Aut_cprod_full _ defZU); rewrite ?cprodE //; last first.
+ by apply: IHs; rewrite -?defZU ?defZE.
+have oZE: #|'Z(E)| = p by rewrite defZE.
+have [p2 | odd_p] := even_prime p_pr.
+ suffices <-: restr_perm 'Z(E) @* Aut E = Aut 'Z(E) by exact: Aut_in_isog.
+ apply/eqP; rewrite eqEcard restr_perm_Aut ?center_sub //=.
+ by rewrite card_Aut_cyclic ?prime_cyclic ?oZE // {1}p2 cardG_gt0.
+have pE: p.-group E by rewrite /pgroup oE pnat_exp pnat_id.
+have nZE: E \subset 'N('Z(E)) by rewrite normal_norm ?center_normal.
+have esE: extraspecial E := card_p3group_extraspecial p_pr oE oZE.
+have [[defPhiE defE'] prZ] := esE.
+have{defPhiE} sEpZ x: x \in E -> (x ^+ p)%g \in 'Z(E).
+ move=> Ex; rewrite -defPhiE (Phi_joing pE) mem_gen // inE orbC.
+ by rewrite (Mho_p_elt 1) // (mem_p_elt pE).
+have ltZE: 'Z(E) \proper E by rewrite properEcard subsetIl oZE oE (ltn_exp2l 1).
+have [x [Ex notZx oxp]]: exists x, [/\ x \in E, x \notin 'Z(E) & #[x] %| p]%N.
+ have [_ [x Ex notZx]] := properP ltZE.
+ case: (prime_subgroupVti <[x ^+ p]> prZ) => [sZxp | ]; last first.
+ move/eqP; rewrite (setIidPl _) ?cycle_subG ?sEpZ //.
+ by rewrite cycle_eq1 -order_dvdn; exists x.
+ have [y Ey notxy]: exists2 y, y \in E & y \notin <[x]>.
+ apply/subsetPn; apply: contra (extraspecial_nonabelian esE) => sEx.
+ by rewrite (abelianS sEx) ?cycle_abelian.
+ have: (y ^+ p)%g \in <[x ^+ p]> by rewrite (subsetP sZxp) ?sEpZ.
+ case/cycleP=> i def_yp; set xi := (x ^- i)%g.
+ have Exi: xi \in E by rewrite groupV groupX.
+ exists (y * xi)%g; split; first by rewrite groupM.
+ have sxpx: <[x ^+ p]> \subset <[x]> by rewrite cycle_subG mem_cycle.
+ apply: contra notxy; move/(subsetP (subset_trans sZxp sxpx)).
+ by rewrite groupMr // groupV mem_cycle.
+ pose z := [~ xi, y]; have Zz: z \in 'Z(E) by rewrite -defE' mem_commg.
+ case: (setIP Zz) => _; move/centP=> cEz.
+ rewrite order_dvdn expMg_Rmul; try by apply: commute_sym; apply: cEz.
+ rewrite def_yp expgVn -!expgM mulnC mulgV mul1g -order_dvdn.
+ by rewrite (dvdn_trans (order_dvdG Zz)) //= oZE bin2odd // dvdn_mulr.
+have{oxp} ox: #[x] = p.
+ apply/eqP; case/primeP: p_pr => _ dvd_p; case/orP: (dvd_p _ oxp) => //.
+ by rewrite order_eq1; case: eqP notZx => // ->; rewrite group1.
+have [y Ey not_cxy]: exists2 y, y \in E & y \notin 'C[x].
+ by apply/subsetPn; rewrite sub_cent1; rewrite inE Ex in notZx.
+have notZy: y \notin 'Z(E).
+ apply: contra not_cxy; rewrite inE Ey; apply: subsetP.
+ by rewrite -cent_set1 centS ?sub1set.
+pose K := 'C_E[y]; have maxK: maximal K E by exact: cent1_extraspecial_maximal.
+have nsKE: K <| E := p_maximal_normal pE maxK; have [sKE nKE] := andP nsKE.
+have oK: #|K| = (p ^ 2)%N.
+ by rewrite -(divg_indexS sKE) oE (p_maximal_index pE) ?mulKn.
+have cKK: abelian K := card_p2group_abelian p_pr oK.
+have sZK: 'Z(E) \subset K by rewrite setIS // -cent_set1 centS ?sub1set.
+have defE: K ><| <[x]> = E.
+ have notKx: x \notin K by rewrite inE Ex cent1C.
+ rewrite sdprodE ?(mulg_normal_maximal nsKE) ?cycle_subG ?(subsetP nKE) //.
+ by rewrite setIC prime_TIg -?orderE ?ox ?cycle_subG.
+have /cyclicP[z defZ]: cyclic 'Z(E) by rewrite prime_cyclic ?oZE.
+apply/(Aut_sub_fullP (center_sub E)); rewrite /= defZ => g injg gZ.
+pose k := invm (injm_Zp_unitm z) (aut injg gZ).
+have fM: {in K &, {morph expgn^~ (val k): u v / u * v}}.
+ by move=> u v Ku Kv; rewrite /= expgMn // /commute (centsP cKK).
+pose f := Morphism fM; have fK: f @* K = K.
+ apply/setP=> u; rewrite morphimEdom.
+ apply/imsetP/idP=> [[v Kv ->] | Ku]; first exact: groupX.
+ exists (u ^+ expg_invn K (val k)); first exact: groupX.
+ rewrite /f /= expgAC expgK // oK coprime_expl // -unitZpE //.
+ by case: (k) => /=; rewrite orderE -defZ oZE => j; rewrite natr_Zp.
+have fMact: {in K & <[x]>, morph_act 'J 'J f (idm <[x]>)}.
+ by move=> u v _ _; rewrite /= conjXg.
+exists (sdprodm_morphism defE fMact).
+rewrite im_sdprodm injm_sdprodm injm_idm -card_im_injm im_idm fK.
+have [_ -> _ ->] := sdprodP defE; rewrite !eqxx; split=> //= u Zu.
+rewrite sdprodmEl ?(subsetP sZK) ?defZ // -(autE injg gZ Zu).
+rewrite -[aut _ _](invmK (injm_Zp_unitm z)); first by rewrite permE Zu.
+by rewrite im_Zp_unitm Aut_aut.
+Qed.
+
+(* These are the parts of Aschbacher (23.12) and exercise 8.5 that are later *)
+(* used in Aschbacher (34.9), which itself replaces the informal discussion *)
+(* quoted from Gorenstein in the proof of B & G, Theorem 2.5. *)
+Lemma center_aut_extraspecial k : coprime k p ->
+ exists2 f, f \in Aut S & forall z, z \in 'Z(S) -> f z = (z ^+ k)%g.
+Proof.
+have /cyclicP[z defZ]: cyclic 'Z(S) by rewrite prime_cyclic ?oZ.
+have oz: #[z] = p by rewrite orderE -defZ.
+rewrite coprime_sym -unitZpE ?prime_gt1 // -oz => u_k.
+pose g := Zp_unitm (FinRing.unit 'Z_#[z] u_k).
+have AutZg: g \in Aut 'Z(S) by rewrite defZ -im_Zp_unitm mem_morphim ?inE.
+have ZSfull := Aut_sub_fullP (center_sub S) Aut_extraspecial_full.
+have [f [injf fS fZ]] := ZSfull _ (injm_autm AutZg) (im_autm AutZg).
+exists (aut injf fS) => [|u Zu]; first exact: Aut_aut.
+have [Su _] := setIP Zu; have z_u: u \in <[z]> by rewrite -defZ.
+by rewrite autE // fZ //= autmE permE /= z_u /cyclem expg_znat.
+Qed.
+
+End StructureCorollaries.
+
+End Extraspecial.
+
+Section SCN.
+
+Variables (gT : finGroupType) (p : nat) (G : {group gT}).
+Implicit Types A Z H : {group gT}.
+
+Lemma SCN_P A : reflect (A <| G /\ 'C_G(A) = A) (A \in 'SCN(G)).
+Proof. by apply: (iffP setIdP) => [] [->]; move/eqP. Qed.
+
+Lemma SCN_abelian A : A \in 'SCN(G) -> abelian A.
+Proof. by case/SCN_P=> _ defA; rewrite /abelian -{1}defA subsetIr. Qed.
+
+Lemma exponent_Ohm1_class2 H :
+ odd p -> p.-group H -> nil_class H <= 2 -> exponent 'Ohm_1(H) %| p.
+Proof.
+move=> odd_p pH; rewrite nil_class2 => sH'Z; apply/exponentP=> x /=.
+rewrite (OhmE 1 pH) expn1 gen_set_id => {x} [/LdivP[] //|].
+apply/group_setP; split=> [|x y]; first by rewrite !inE group1 expg1n //=.
+case/LdivP=> Hx xp1 /LdivP[Hy yp1]; rewrite !inE groupM //=.
+have [_ czH]: [~ y, x] \in H /\ centralises [~ y, x] H.
+ by apply/centerP; rewrite (subsetP sH'Z) ?mem_commg.
+rewrite expMg_Rmul ?xp1 ?yp1 /commute ?czH //= !mul1g.
+by rewrite bin2odd // -commXXg ?yp1 /commute ?czH // comm1g.
+Qed.
+
+(* SCN_max and max_SCN cover Aschbacher 23.15(1) *)
+Lemma SCN_max A : A \in 'SCN(G) -> [max A | A <| G & abelian A].
+Proof.
+case/SCN_P => nAG scA; apply/maxgroupP; split=> [|H].
+ by rewrite nAG /abelian -{1}scA subsetIr.
+do 2![case/andP]=> sHG _ abelH sAH; apply/eqP.
+by rewrite eqEsubset sAH -scA subsetI sHG centsC (subset_trans sAH).
+Qed.
+
+Lemma max_SCN A :
+ p.-group G -> [max A | A <| G & abelian A] -> A \in 'SCN(G).
+Proof.
+move/pgroup_nil=> nilG; rewrite /abelian.
+case/maxgroupP=> /andP[nsAG abelA] maxA; have [sAG nAG] := andP nsAG.
+rewrite inE nsAG eqEsubset /= andbC subsetI abelA normal_sub //=.
+rewrite -quotient_sub1; last by rewrite subIset 1?normal_norm.
+apply/trivgP; apply: (TI_center_nil (quotient_nil A nilG)).
+ by rewrite quotient_normal // /normal subsetIl normsI ?normG ?norms_cent.
+apply/trivgP/subsetP=> _ /setIP[/morphimP[x Nx /setIP[_ Cx]] ->].
+rewrite -cycle_subG in Cx => /setIP[GAx CAx].
+have{CAx GAx}: <[coset A x]> <| G / A.
+ by rewrite /normal cycle_subG GAx cents_norm // centsC cycle_subG.
+case/(inv_quotientN nsAG)=> B /= defB sAB nBG.
+rewrite -cycle_subG defB (maxA B) ?trivg_quotient // nBG.
+have{defB} defB : B :=: A * <[x]>.
+ rewrite -quotientK ?cycle_subG ?quotient_cycle // defB quotientGK //.
+ exact: normalS (normal_sub nBG) nsAG.
+apply/setIidPl; rewrite ?defB -[_ :&: _]center_prod //=.
+rewrite /center !(setIidPl _) //; exact: cycle_abelian.
+Qed.
+
+(* The two other assertions of Aschbacher 23.15 state properties of the *)
+(* normal series 1 <| Z = 'Ohm_1(A) <| A with A \in 'SCN(G). *)
+
+Section SCNseries.
+
+Variables A : {group gT}.
+Hypothesis SCN_A : A \in 'SCN(G).
+
+Let Z := 'Ohm_1(A).
+Let cAA := SCN_abelian SCN_A.
+Let sZA: Z \subset A := Ohm_sub 1 A.
+Let nZA : A \subset 'N(Z) := sub_abelian_norm cAA sZA.
+
+(* This is Aschbacher 23.15(2). *)
+Lemma der1_stab_Ohm1_SCN_series : ('C(Z) :&: 'C_G(A / Z | 'Q))^`(1) \subset A.
+Proof.
+case/SCN_P: SCN_A => /andP[sAG nAG] {4} <-.
+rewrite subsetI {1}setICA comm_subG ?subsetIl //= gen_subG.
+apply/subsetP=> w /imset2P[u v].
+rewrite -groupV -(groupV _ v) /= astabQR //= -/Z !inE groupV.
+case/and4P=> cZu _ _ sRuZ /and4P[cZv' _ _ sRvZ] ->{w}.
+apply/centP=> a Aa; rewrite /commute -!mulgA (commgCV v) (mulgA u).
+rewrite (centP cZu); last by rewrite (subsetP sRvZ) ?mem_commg ?set11 ?groupV.
+rewrite 2!(mulgA v^-1) mulKVg 4!mulgA invgK (commgC u^-1) mulgA.
+rewrite -(mulgA _ _ v^-1) -(centP cZv') ?(subsetP sRuZ) ?mem_commg ?set11//.
+by rewrite -!mulgA invgK mulKVg !mulKg.
+Qed.
+
+(* This is Aschbacher 23.15(3); note that this proof does not depend on the *)
+(* maximality of A. *)
+Lemma Ohm1_stab_Ohm1_SCN_series :
+ odd p -> p.-group G -> 'Ohm_1('C_G(Z)) \subset 'C_G(A / Z | 'Q).
+Proof.
+have [-> | ntG] := eqsVneq G 1; first by rewrite !(setIidPl (sub1G _)) Ohm1.
+move=> p_odd pG; have{ntG} [p_pr _ _] := pgroup_pdiv pG ntG.
+case/SCN_P: SCN_A => /andP[sAG nAG] _; have pA := pgroupS sAG pG.
+have pCGZ : p.-group 'C_G(Z) by rewrite (pgroupS _ pG) // subsetIl.
+rewrite {pCGZ}(OhmE 1 pCGZ) gen_subG; apply/subsetP=> x; rewrite 3!inE -andbA.
+rewrite -!cycle_subG => /and3P[sXG cZX xp1] /=; have cXX := cycle_abelian x.
+have nZX := cents_norm cZX; have{nAG} nAX := subset_trans sXG nAG.
+pose XA := <[x]> <*> A; pose C := 'C(<[x]> / Z | 'Q); pose CA := A :&: C.
+pose Y := <[x]> <*> CA; pose W := 'Ohm_1(Y).
+have sXC: <[x]> \subset C by rewrite sub_astabQ nZX (quotient_cents _ cXX).
+have defY : Y = <[x]> * CA by rewrite -norm_joinEl // normsI ?nAX ?normsG.
+have{nAX} defXA: XA = <[x]> * A := norm_joinEl nAX.
+suffices{sXC}: XA \subset Y.
+ rewrite subsetI sXG /= sub_astabQ nZX centsC defY group_modl //= -/Z -/C.
+ by rewrite subsetI sub_astabQ defXA quotientMl //= !mulG_subG; case/and4P.
+have sZCA: Z \subset CA by rewrite subsetI sZA [C]astabQ sub_cosetpre.
+have cZCA: CA \subset 'C(Z) by rewrite subIset 1?(sub_abelian_cent2 cAA).
+have sZY: Z \subset Y by rewrite (subset_trans sZCA) ?joing_subr.
+have{cZCA cZX} cZY: Y \subset 'C(Z) by rewrite join_subG cZX.
+have{cXX nZX} sY'Z : Y^`(1) \subset Z.
+ rewrite der1_min ?cents_norm //= -/Y defY quotientMl // abelianM /= -/Z -/CA.
+ rewrite !quotient_abelian // ?(abelianS _ cAA) ?subsetIl //=.
+ by rewrite /= quotientGI ?Ohm_sub // quotient_astabQ subsetIr.
+have{sY'Z cZY} nil_classY: nil_class Y <= 2.
+ by rewrite nil_class2 (subset_trans sY'Z ) // subsetI sZY centsC.
+have pY: p.-group Y by rewrite (pgroupS _ pG) // join_subG sXG subIset ?sAG.
+have sXW: <[x]> \subset W.
+ by rewrite [W](OhmE 1 pY) ?genS // sub1set !inE -cycle_subG joing_subl.
+have{nil_classY pY sXW sZY sZCA} defW: W = <[x]> * Z.
+ rewrite -[W](setIidPr (Ohm_sub _ _)) /= -/Y {1}defY -group_modl //= -/Y -/W.
+ congr (_ * _); apply/eqP; rewrite eqEsubset {1}[Z](OhmE 1 pA).
+ rewrite subsetI setIAC subIset //; first by rewrite sZCA -[Z]Ohm_id OhmS.
+ rewrite sub_gen ?setIS //; apply/subsetP=> w Ww; rewrite inE.
+ by apply/eqP; apply: exponentP w Ww; exact: exponent_Ohm1_class2.
+have{sXG sAG} sXAG: XA \subset G by rewrite join_subG sXG.
+have{sXAG} nilXA: nilpotent XA := nilpotentS sXAG (pgroup_nil pG).
+have sYXA: Y \subset XA by rewrite defY defXA mulgS ?subsetIl.
+rewrite -[Y](nilpotent_sub_norm nilXA) {nilXA sYXA}//= -/Y -/XA.
+apply: subset_trans (setIS _ (char_norm_trans (Ohm_char 1 _) (subxx _))) _.
+rewrite {XA}defXA -group_modl ?normsG /= -/W ?{W}defW ?mulG_subl //.
+rewrite {Y}defY mulgS // subsetI subsetIl {CA C}sub_astabQ subIset ?nZA //= -/Z.
+rewrite (subset_trans (quotient_subnorm _ _ _)) //= quotientMidr /= -/Z.
+rewrite -quotient_sub1 ?subIset ?cent_norm ?orbT //.
+rewrite (subset_trans (quotientI _ _ _)) ?coprime_TIg //.
+rewrite (@pnat_coprime p) // -/(pgroup p _) ?quotient_pgroup {pA}//=.
+rewrite -(setIidPr (cent_sub _)) [pnat _ _]p'group_quotient_cent_prime //.
+by rewrite (dvdn_trans (dvdn_quotient _ _)) ?order_dvdn.
+Qed.
+
+End SCNseries.
+
+(* This is Aschbacher 23.16. *)
+Lemma Ohm1_cent_max_normal_abelem Z :
+ odd p -> p.-group G -> [max Z | Z <| G & p.-abelem Z] -> 'Ohm_1('C_G(Z)) = Z.
+Proof.
+move=> p_odd pG; set X := 'Ohm_1('C_G(Z)).
+case/maxgroupP=> /andP[nsZG abelZ] maxZ.
+have [sZG nZG] := andP nsZG; have [_ cZZ expZp] := and3P abelZ.
+have{nZG} nsXG: X <| G.
+ apply: (char_normal_trans (Ohm_char 1 'C_G(Z))).
+ by rewrite /normal subsetIl normsI ?normG ?norms_cent.
+have cZX : X \subset 'C(Z) := subset_trans (Ohm_sub _ _) (subsetIr _ _).
+have{sZG expZp} sZX: Z \subset X.
+ rewrite [X](OhmE 1 (pgroupS _ pG)) ?subsetIl ?sub_gen //.
+ apply/subsetP=> x Zx; rewrite !inE ?(subsetP sZG) ?(subsetP cZZ) //=.
+ by rewrite (exponentP expZp).
+suffices{sZX} expXp: (exponent X %| p).
+ apply/eqP; rewrite eqEsubset sZX andbT -quotient_sub1 ?cents_norm //= -/X.
+ have pGq: p.-group (G / Z) by rewrite quotient_pgroup.
+ rewrite (TI_center_nil (pgroup_nil pGq)) ?quotient_normal //= -/X setIC.
+ apply/eqP/trivgPn=> [[Zd]]; rewrite inE -!cycle_subG -cycle_eq1 -subG1 /= -/X.
+ case/andP=> /sub_center_normal nsZdG.
+ have{nsZdG} [D defD sZD nsDG] := inv_quotientN nsZG nsZdG; rewrite defD.
+ have sDG := normal_sub nsDG; have nsZD := normalS sZD sDG nsZG.
+ rewrite quotientSGK ?quotient_sub1 ?normal_norm //= -/X => sDX; case/negP.
+ rewrite (maxZ D) // nsDG andbA (pgroupS sDG) ?(dvdn_trans (exponentS sDX)) //.
+ have sZZD: Z \subset 'Z(D) by rewrite subsetI sZD centsC (subset_trans sDX).
+ by rewrite (cyclic_factor_abelian sZZD) //= -defD cycle_cyclic.
+pose normal_abelian := [pred A : {group gT} | A <| G & abelian A].
+have{nsZG cZZ} normal_abelian_Z : normal_abelian Z by exact/andP.
+have{normal_abelian_Z} [A maxA sZA] := maxgroup_exists normal_abelian_Z.
+have SCN_A : A \in 'SCN(G) by apply: max_SCN pG maxA.
+move/maxgroupp: maxA => /andP[nsAG cAA] {normal_abelian}.
+have pA := pgroupS (normal_sub nsAG) pG.
+have{abelZ maxZ nsAG cAA sZA} defA1: 'Ohm_1(A) = Z.
+ apply: maxZ; last by rewrite -(Ohm1_id abelZ) OhmS.
+ by rewrite Ohm1_abelem ?(char_normal_trans (Ohm_char _ _) nsAG).
+have{SCN_A} sX'A: X^`(1) \subset A.
+ have sX_CWA1 : X \subset 'C('Ohm_1(A)) :&: 'C_G(A / 'Ohm_1(A) | 'Q).
+ rewrite subsetI /X -defA1 (Ohm1_stab_Ohm1_SCN_series _ p_odd) //= andbT.
+ exact: subset_trans (Ohm_sub _ _) (subsetIr _ _).
+ by apply: subset_trans (der1_stab_Ohm1_SCN_series SCN_A); rewrite commgSS.
+pose genXp := [pred U : {group gT} | 'Ohm_1(U) == U & ~~ (exponent U %| p)].
+apply/idPn=> expXp'; have genXp_X: genXp [group of X] by rewrite /= Ohm_id eqxx.
+have{genXp_X expXp'} [U] := mingroup_exists genXp_X; case/mingroupP; case/andP.
+move/eqP=> defU1 expUp' minU sUX; case/negP: expUp'.
+have{nsXG} pU := pgroupS (subset_trans sUX (normal_sub nsXG)) pG.
+case gsetU1: (group_set 'Ldiv_p(U)).
+ by rewrite -defU1 (OhmE 1 pU) gen_set_id // -sub_LdivT subsetIr.
+move: gsetU1; rewrite /group_set 2!inE group1 expg1n eqxx; case/subsetPn=> xy.
+case/imset2P=> x y; rewrite !inE => /andP[Ux xp1] /andP[Uy yp1] ->{xy}.
+rewrite groupM //= => nt_xyp; pose XY := <[x]> <*> <[y]>.
+have{yp1 nt_xyp} defXY: XY = U.
+ have sXY_U: XY \subset U by rewrite join_subG !cycle_subG Ux Uy.
+ rewrite [XY]minU //= eqEsubset Ohm_sub (OhmE 1 (pgroupS _ pU)) //.
+ rewrite /= joing_idl joing_idr genS; last first.
+ by rewrite subsetI subset_gen subUset !sub1set !inE xp1 yp1.
+ apply: contra nt_xyp => /exponentP-> //.
+ by rewrite groupMl mem_gen // (set21, set22).
+have: <[x]> <|<| U by rewrite nilpotent_subnormal ?(pgroup_nil pU) ?cycle_subG.
+case/subnormalEsupport=> [defU | /=].
+ by apply: dvdn_trans (exponent_dvdn U) _; rewrite -defU order_dvdn.
+set V := <<class_support <[x]> U>>; case/andP=> sVU ltVU.
+have{genXp minU xp1 sVU ltVU} expVp: exponent V %| p.
+ apply: contraR ltVU => expVp'; rewrite [V]minU //= expVp' eqEsubset Ohm_sub.
+ rewrite (OhmE 1 (pgroupS sVU pU)) genS //= subsetI subset_gen class_supportEr.
+ apply/bigcupsP=> z _; apply/subsetP=> v Vv.
+ by rewrite inE -order_dvdn (dvdn_trans (order_dvdG Vv)) // cardJg order_dvdn.
+have{A pA defA1 sX'A V expVp} Zxy: [~ x, y] \in Z.
+ rewrite -defA1 (OhmE 1 pA) mem_gen // !inE (exponentP expVp).
+ by rewrite (subsetP sX'A) //= mem_commg ?(subsetP sUX).
+ by rewrite groupMl -1?[x^-1]conjg1 mem_gen // mem_imset2 // ?groupV cycle_id.
+have{Zxy sUX cZX} cXYxy: [~ x, y] \in 'C(XY).
+ by rewrite centsC in cZX; rewrite defXY (subsetP (centS sUX)) ?(subsetP cZX).
+rewrite -defU1 exponent_Ohm1_class2 // nil_class2 -defXY der1_joing_cycles //.
+by rewrite subsetI {1}defXY !cycle_subG groupR.
+Qed.
+
+Lemma critical_class2 H : critical H G -> nil_class H <= 2.
+Proof.
+case=> [chH _ sRZ _].
+by rewrite nil_class2 (subset_trans _ sRZ) ?commSg // char_sub.
+Qed.
+
+(* This proof of the Thompson critical lemma is adapted from Aschbacher 23.6 *)
+Lemma Thompson_critical : p.-group G -> {K : {group gT} | critical K G}.
+Proof.
+move=> pG; pose qcr A := (A \char G) && ('Phi(A) :|: [~: G, A] \subset 'Z(A)).
+have [|K]:= @maxgroup_exists _ qcr 1 _.
+ by rewrite /qcr char1 center1 commG1 subUset Phi_sub subxx.
+case/maxgroupP; rewrite {}/qcr subUset => /and3P[chK sPhiZ sRZ] maxK _.
+have sKG := char_sub chK; have nKG := char_normal chK.
+exists K; split=> //; apply/eqP; rewrite eqEsubset andbC setSI //=.
+have chZ: 'Z(K) \char G by [exact: subcent_char]; have nZG := char_norm chZ.
+have chC: 'C_G(K) \char G by exact: subcent_char (char_refl G) chK.
+rewrite -quotient_sub1; last by rewrite subIset // char_norm.
+apply/trivgP; apply: (TI_center_nil (quotient_nil _ (pgroup_nil pG))).
+ rewrite quotient_normal // /normal subsetIl normsI ?normG ?norms_cent //.
+ exact: char_norm.
+apply: TI_Ohm1; apply/trivgP; rewrite -trivg_quotient -sub_cosetpre_quo //.
+rewrite morphpreI quotientGK /=; last first.
+ by apply: normalS (char_normal chZ); rewrite ?subsetIl ?setSI.
+set X := _ :&: _; pose gX := [group of X].
+have sXG: X \subset G by rewrite subIset ?subsetIl.
+have cXK: K \subset 'C(gX) by rewrite centsC 2?subIset // subxx orbT.
+rewrite subsetI centsC cXK andbT -(mul1g K) -mulSG mul1g -(cent_joinEr cXK).
+rewrite [_ <*> K]maxK ?joing_subr //= andbC (cent_joinEr cXK).
+rewrite -center_prod // (subset_trans _ (mulG_subr _ _)).
+ rewrite charM 1?charI ?(char_from_quotient (normal_cosetpre _)) //.
+ by rewrite cosetpreK (char_trans _ (center_char _)) ?Ohm_char.
+rewrite (@Phi_mulg p) ?(pgroupS _ pG) // subUset commGC commMG; last first.
+ by rewrite normsR ?(normsG sKG) // cents_norm // centsC.
+rewrite !mul_subG 1?commGC //.
+ apply: subset_trans (commgS _ (subsetIr _ _)) _.
+ rewrite -quotient_cents2 ?subsetIl // centsC // cosetpreK //.
+ by rewrite (subset_trans (Ohm_sub _ _)) // subsetIr.
+have nZX := subset_trans sXG nZG; have pX : p.-group gX by exact: pgroupS pG.
+rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) //=.
+have pXZ: p.-group (gX / 'Z(K)) by exact: morphim_pgroup.
+rewrite (quotient_Phi pX nZX) subG1 (trivg_Phi pXZ).
+apply: (abelemS (quotientS _ (subsetIr _ _))); rewrite /= cosetpreK /=.
+have pZ: p.-group 'Z(G / 'Z(K)).
+ by rewrite (pgroupS (center_sub _)) ?morphim_pgroup.
+by rewrite Ohm1_abelem ?center_abelian.
+Qed.
+
+Lemma critical_p_stab_Aut H :
+ critical H G -> p.-group G -> p.-group 'C(H | [Aut G]).
+Proof.
+move=> [chH sPhiZ sRZ eqCZ] pG; have sHG := char_sub chH.
+pose G' := (sdpair1 [Aut G] @* G)%G; pose H' := (sdpair1 [Aut G] @* H)%G.
+apply/pgroupP=> q pr_q; case/Cauchy=>//= f cHF; move: (cHF);rewrite astab_ract.
+case/setIP=> Af cHFP ofq; rewrite -cycle_subG in cHF; apply: (pgroupP pG) => //.
+pose F' := (sdpair2 [Aut G] @* <[f]>)%G.
+have trHF: [~: H', F'] = 1.
+ apply/trivgP; rewrite gen_subG; apply/subsetP=> u; case/imset2P=> x' a'.
+ case/morphimP=> x Gx Hx ->; case/morphimP=> a Aa Fa -> -> {u x' a'}.
+ by rewrite inE commgEl -sdpair_act ?(astab_act (subsetP cHF _ Fa) Hx) ?mulVg.
+have sGH_H: [~: G', H'] \subset H'.
+ by rewrite -morphimR ?(char_sub chH) // morphimS // commg_subr char_norm.
+have{trHF sGH_H} trFGH: [~: F', G', H'] = 1.
+ apply: three_subgroup; last by rewrite trHF comm1G.
+ by apply/trivgP; rewrite -trHF commSg.
+apply/negP=> qG; case: (qG); rewrite -ofq.
+suffices ->: f = 1 by rewrite order1 dvd1n.
+apply/permP=> x; rewrite perm1; case Gx: (x \in G); last first.
+ by apply: out_perm (negbT Gx); case/setIdP: Af.
+have Gfx: f x \in G by rewrite -(im_autm Af) -{1}(autmE Af) mem_morphim.
+pose y := x^-1 * f x; have Gy: y \in G by rewrite groupMl ?groupV.
+have [inj1 inj2] := (injm_sdpair1 [Aut G], injm_sdpair2 [Aut G]).
+have Hy: y \in H.
+ rewrite (subsetP (center_sub H)) // -eqCZ -cycle_subG.
+ rewrite -(injmSK inj1) ?cycle_subG // injm_subcent // subsetI.
+ rewrite morphimS ?morphim_cycle ?cycle_subG //=.
+ suffices: sdpair1 [Aut G] y \in [~: G', F'].
+ by rewrite commGC; apply: subsetP; exact/commG1P.
+ rewrite morphM ?groupV ?morphV //= sdpair_act // -commgEl.
+ by rewrite mem_commg ?mem_morphim ?cycle_id.
+have fy: f y = y := astabP cHFP _ Hy.
+have: (f ^+ q) x = x * y ^+ q.
+ elim: (q) => [|i IHi]; first by rewrite perm1 mulg1.
+ rewrite expgSr permM {}IHi -(autmE Af) morphM ?morphX ?groupX //= autmE.
+ by rewrite fy expgS mulgA mulKVg.
+move/eqP; rewrite -{1}ofq expg_order perm1 eq_mulVg1 mulKg -order_dvdn.
+case/primeP: pr_q => _ pr_q /pr_q; rewrite order_eq1 -eq_mulVg1.
+by case: eqP => //= _ /eqP oyq; case: qG; rewrite -oyq order_dvdG.
+Qed.
+
+End SCN.
+
+Implicit Arguments SCN_P [gT G A]. \ No newline at end of file
diff --git a/mathcomp/solvable/nilpotent.v b/mathcomp/solvable/nilpotent.v
new file mode 100644
index 0000000..5b271de
--- /dev/null
+++ b/mathcomp/solvable/nilpotent.v
@@ -0,0 +1,755 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path fintype div.
+Require Import bigop prime finset fingroup morphism automorphism quotient.
+Require Import commutator gproduct gfunctor center gseries cyclic.
+
+(******************************************************************************)
+(* This file defines nilpotent and solvable groups, and give some of their *)
+(* elementary properties; more will be added later (e.g., the nilpotence of *)
+(* p-groups in sylow.v, or the fact that minimal normal subgroups of solvable *)
+(* groups are elementary abelian in maximal.v). This file defines: *)
+(* nilpotent G == G is nilpotent, i.e., [~: H, G] is a proper subgroup of H *)
+(* for all nontrivial H <| G. *)
+(* solvable G == G is solvable, i.e., H^`(1) is a proper subgroup of H for *)
+(* all nontrivial subgroups H of G. *)
+(* 'L_n(G) == the nth term of the lower central series, namely *)
+(* [~: G, ..., G] (n Gs) if n > 0, with 'L_0(G) = G. *)
+(* G is nilpotent iff 'L_n(G) = 1 for some n. *)
+(* 'Z_n(G) == the nth term of the upper central series, i.e., *)
+(* with 'Z_0(G) = 1, 'Z_n.+1(G) / 'Z_n(G) = 'Z(G / 'Z_n(G)). *)
+(* nil_class G == the nilpotence class of G, i.e., the least n such that *)
+(* 'L_n.+1(G) = 1 (or, equivalently, 'Z_n(G) = G), if G is *)
+(* nilpotent; we take nil_class G = #|G| when G is not *)
+(* nilpotent, so nil_class G < #|G| iff G is nilpotent. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section SeriesDefs.
+
+Variables (n : nat) (gT : finGroupType) (A : {set gT}).
+
+Definition lower_central_at_rec := iter n (fun B => [~: B, A]) A.
+
+Definition upper_central_at_rec := iter n (fun B => coset B @*^-1 'Z(A / B)) 1.
+
+End SeriesDefs.
+
+(* By convention, the lower central series starts at 1 while the upper series *)
+(* starts at 0 (sic). *)
+Definition lower_central_at n := lower_central_at_rec n.-1.
+
+(* Note: 'nosimpl' MUST be used outside of a section -- the end of section *)
+(* "cooking" destroys it. *)
+Definition upper_central_at := nosimpl upper_central_at_rec.
+
+Arguments Scope lower_central_at [nat_scope _ group_scope].
+Arguments Scope upper_central_at [nat_scope _ group_scope].
+
+Notation "''L_' n ( G )" := (lower_central_at n G)
+ (at level 8, n at level 2, format "''L_' n ( G )") : group_scope.
+
+Notation "''Z_' n ( G )" := (upper_central_at n G)
+ (at level 8, n at level 2, format "''Z_' n ( G )") : group_scope.
+
+Section PropertiesDefs.
+
+Variables (gT : finGroupType) (A : {set gT}).
+
+Definition nilpotent :=
+ [forall (G : {group gT} | G \subset A :&: [~: G, A]), G :==: 1].
+
+Definition nil_class := index 1 (mkseq (fun n => 'L_n.+1(A)) #|A|).
+
+Definition solvable :=
+ [forall (G : {group gT} | G \subset A :&: [~: G, G]), G :==: 1].
+
+End PropertiesDefs.
+
+Arguments Scope nilpotent [_ group_scope].
+Arguments Scope nil_class [_ group_scope].
+Arguments Scope solvable [_ group_scope].
+Prenex Implicits nil_class nilpotent solvable.
+
+Section NilpotentProps.
+
+Variable gT: finGroupType.
+Implicit Types (A B : {set gT}) (G H : {group gT}).
+
+Lemma nilpotent1 : nilpotent [1 gT].
+Proof. by apply/forall_inP=> H; rewrite commG1 setIid -subG1. Qed.
+
+Lemma nilpotentS A B : B \subset A -> nilpotent A -> nilpotent B.
+Proof.
+move=> sBA nilA; apply/forall_inP=> H sHR.
+have:= forallP nilA H; rewrite (subset_trans sHR) //.
+by apply: subset_trans (setIS _ _) (setSI _ _); rewrite ?commgS.
+Qed.
+
+Lemma nil_comm_properl G H A :
+ nilpotent G -> H \subset G -> H :!=: 1 -> A \subset 'N_G(H) ->
+ [~: H, A] \proper H.
+Proof.
+move=> nilG sHG ntH; rewrite subsetI properE; case/andP=> sAG nHA.
+rewrite (subset_trans (commgS H (subset_gen A))) ?commg_subl ?gen_subG //.
+apply: contra ntH => sHR; have:= forallP nilG H; rewrite subsetI sHG.
+by rewrite (subset_trans sHR) ?commgS.
+Qed.
+
+Lemma nil_comm_properr G A H :
+ nilpotent G -> H \subset G -> H :!=: 1 -> A \subset 'N_G(H) ->
+ [~: A, H] \proper H.
+Proof. by rewrite commGC; apply: nil_comm_properl. Qed.
+
+Lemma centrals_nil (s : seq {group gT}) G :
+ G.-central.-series 1%G s -> last 1%G s = G -> nilpotent G.
+Proof.
+move=> cGs defG; apply/forall_inP=> H /subsetIP[sHG sHR].
+move: sHG; rewrite -{}defG -subG1 -[1]/(gval 1%G).
+elim: s 1%G cGs => //= L s IHs K /andP[/and3P[sRK sKL sLG] /IHs sHL] sHs.
+exact: subset_trans sHR (subset_trans (commSg _ (sHL sHs)) sRK).
+Qed.
+
+End NilpotentProps.
+
+Section LowerCentral.
+
+Variable gT : finGroupType.
+Implicit Types (A B : {set gT}) (G H : {group gT}).
+
+Lemma lcn0 A : 'L_0(A) = A. Proof. by []. Qed.
+Lemma lcn1 A : 'L_1(A) = A. Proof. by []. Qed.
+Lemma lcnSn n A : 'L_n.+2(A) = [~: 'L_n.+1(A), A]. Proof. by []. Qed.
+Lemma lcnSnS n G : [~: 'L_n(G), G] \subset 'L_n.+1(G).
+Proof. by case: n => //; exact: der1_subG. Qed.
+Lemma lcnE n A : 'L_n.+1(A) = lower_central_at_rec n A.
+Proof. by []. Qed.
+Lemma lcn2 A : 'L_2(A) = A^`(1). Proof. by []. Qed.
+
+Lemma lcn_group_set n G : group_set 'L_n(G).
+Proof. by case: n => [|[|n]]; exact: groupP. Qed.
+
+Canonical lower_central_at_group n G := Group (lcn_group_set n G).
+
+Lemma lcn_char n G : 'L_n(G) \char G.
+Proof.
+by case: n => [|n]; last elim: n => [|n IHn]; rewrite ?lcnSn ?charR ?char_refl.
+Qed.
+
+Lemma lcn_normal n G : 'L_n(G) <| G.
+Proof. by apply: char_normal; exact: lcn_char. Qed.
+
+Lemma lcn_sub n G : 'L_n(G) \subset G.
+Proof. by case/andP: (lcn_normal n G). Qed.
+
+Lemma lcn_norm n G : G \subset 'N('L_n(G)).
+Proof. by case/andP: (lcn_normal n G). Qed.
+
+Lemma lcn_subS n G : 'L_n.+1(G) \subset 'L_n(G).
+Proof.
+case: n => // n; rewrite lcnSn commGC commg_subr.
+by case/andP: (lcn_normal n.+1 G).
+Qed.
+
+Lemma lcn_normalS n G : 'L_n.+1(G) <| 'L_n(G).
+Proof. by apply: normalS (lcn_normal _ _); rewrite (lcn_subS, lcn_sub). Qed.
+
+Lemma lcn_central n G : 'L_n(G) / 'L_n.+1(G) \subset 'Z(G / 'L_n.+1(G)).
+Proof.
+case: n => [|n]; first by rewrite trivg_quotient sub1G.
+by rewrite subsetI quotientS ?lcn_sub ?quotient_cents2r.
+Qed.
+
+Lemma lcn_sub_leq m n G : n <= m -> 'L_m(G) \subset 'L_n(G).
+Proof.
+by move/subnK <-; elim: {m}(m - n) => // m; exact: subset_trans (lcn_subS _ _).
+Qed.
+
+Lemma lcnS n A B : A \subset B -> 'L_n(A) \subset 'L_n(B).
+Proof.
+by case: n => // n sAB; elim: n => // n IHn; rewrite !lcnSn genS ?imset2S.
+Qed.
+
+Lemma lcn_cprod n A B G : A \* B = G -> 'L_n(A) \* 'L_n(B) = 'L_n(G).
+Proof.
+case: n => // n /cprodP[[H K -> ->{A B}] defG cHK].
+have sL := subset_trans (lcn_sub _ _); rewrite cprodE ?(centSS _ _ cHK) ?sL //.
+symmetry; elim: n => // n; rewrite lcnSn => ->; rewrite commMG /=; last first.
+ by apply: subset_trans (commg_normr _ _); rewrite sL // -defG mulG_subr.
+rewrite -!(commGC G) -defG -{1}(centC cHK).
+rewrite !commMG ?normsR ?lcn_norm ?cents_norm // 1?centsC //.
+by rewrite -!(commGC 'L__(_)) -!lcnSn !(commG1P _) ?mul1g ?sL // centsC.
+Qed.
+
+Lemma lcn_dprod n A B G : A \x B = G -> 'L_n(A) \x 'L_n(B) = 'L_n(G).
+Proof.
+move=> defG; have [[K H defA defB] _ _ tiAB] := dprodP defG.
+rewrite !dprodEcp // in defG *; first exact: lcn_cprod.
+by rewrite defA defB; apply/trivgP; rewrite -tiAB defA defB setISS ?lcn_sub.
+Qed.
+
+Lemma der_cprod n A B G : A \* B = G -> A^`(n) \* B^`(n) = G^`(n).
+Proof. by move=> defG; elim: n => {defG}// n; apply: (lcn_cprod 2). Qed.
+
+Lemma der_dprod n A B G : A \x B = G -> A^`(n) \x B^`(n) = G^`(n).
+Proof. by move=> defG; elim: n => {defG}// n; apply: (lcn_dprod 2). Qed.
+
+Lemma lcn_bigcprod n I r P (F : I -> {set gT}) G :
+ \big[cprod/1]_(i <- r | P i) F i = G ->
+ \big[cprod/1]_(i <- r | P i) 'L_n(F i) = 'L_n(G).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first exact/esym/trivgP/lcn_sub.
+by rewrite -(lcn_cprod n dG); have [[_ H _ dH]] := cprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma lcn_bigdprod n I r P (F : I -> {set gT}) G :
+ \big[dprod/1]_(i <- r | P i) F i = G ->
+ \big[dprod/1]_(i <- r | P i) 'L_n(F i) = 'L_n(G).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first exact/esym/trivgP/lcn_sub.
+by rewrite -(lcn_dprod n dG); have [[_ H _ dH]] := dprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma der_bigcprod n I r P (F : I -> {set gT}) G :
+ \big[cprod/1]_(i <- r | P i) F i = G ->
+ \big[cprod/1]_(i <- r | P i) (F i)^`(n) = G^`(n).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first by rewrite gF1.
+by rewrite -(der_cprod n dG); have [[_ H _ dH]] := cprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma der_bigdprod n I r P (F : I -> {set gT}) G :
+ \big[dprod/1]_(i <- r | P i) F i = G ->
+ \big[dprod/1]_(i <- r | P i) (F i)^`(n) = G^`(n).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first by rewrite gF1.
+by rewrite -(der_dprod n dG); have [[_ H _ dH]] := dprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma nilpotent_class G : nilpotent G = (nil_class G < #|G|).
+Proof.
+rewrite /nil_class; set s := mkseq _ _.
+transitivity (1 \in s); last by rewrite -index_mem size_mkseq.
+apply/idP/mapP=> {s}/= [nilG | [n _ Ln1]]; last first.
+ apply/forall_inP=> H /subsetIP[sHG sHR].
+ rewrite -subG1 {}Ln1; elim: n => // n IHn.
+ by rewrite (subset_trans sHR) ?commSg.
+pose m := #|G|.-1; exists m; first by rewrite mem_iota /= prednK.
+rewrite ['L__(G)]card_le1_trivg //= -(subnn m).
+elim: {-2}m => [|n]; [by rewrite subn0 prednK | rewrite lcnSn subnS].
+case: (eqsVneq 'L_n.+1(G) 1) => [-> | ntLn]; first by rewrite comm1G cards1.
+case: (m - n) => [|m' /= IHn]; first by rewrite leqNgt cardG_gt1 ntLn.
+rewrite -ltnS (leq_trans (proper_card _) IHn) //.
+by rewrite (nil_comm_properl nilG) ?lcn_sub // subsetI subxx lcn_norm.
+Qed.
+
+Lemma lcn_nil_classP n G :
+ nilpotent G -> reflect ('L_n.+1(G) = 1) (nil_class G <= n).
+Proof.
+rewrite nilpotent_class /nil_class; set s := mkseq _ _.
+set c := index 1 s => lt_c_G; case: leqP => [le_c_n | lt_n_c].
+ have Lc1: nth 1 s c = 1 by rewrite nth_index // -index_mem size_mkseq.
+ by left; apply/trivgP; rewrite -Lc1 nth_mkseq ?lcn_sub_leq.
+right; apply/eqP/negPf; rewrite -(before_find 1 lt_n_c) nth_mkseq //.
+exact: ltn_trans lt_n_c lt_c_G.
+Qed.
+
+Lemma lcnP G : reflect (exists n, 'L_n.+1(G) = 1) (nilpotent G).
+Proof.
+apply: (iffP idP) => [nilG | [n Ln1]].
+ by exists (nil_class G); exact/lcn_nil_classP.
+apply/forall_inP=> H /subsetIP[sHG sHR]; rewrite -subG1 -{}Ln1.
+by elim: n => // n IHn; rewrite (subset_trans sHR) ?commSg.
+Qed.
+
+Lemma abelian_nil G : abelian G -> nilpotent G.
+Proof. by move=> abG; apply/lcnP; exists 1%N; exact/commG1P. Qed.
+
+Lemma nil_class0 G : (nil_class G == 0) = (G :==: 1).
+Proof.
+apply/idP/eqP=> [nilG | ->].
+ by apply/(lcn_nil_classP 0); rewrite ?nilpotent_class (eqP nilG) ?cardG_gt0.
+by rewrite -leqn0; apply/(lcn_nil_classP 0); rewrite ?nilpotent1.
+Qed.
+
+Lemma nil_class1 G : (nil_class G <= 1) = abelian G.
+Proof.
+have [-> | ntG] := eqsVneq G 1.
+ by rewrite abelian1 leq_eqVlt ltnS leqn0 nil_class0 eqxx orbT.
+apply/idP/idP=> cGG.
+ apply/commG1P; apply/(lcn_nil_classP 1); rewrite // nilpotent_class.
+ by rewrite (leq_ltn_trans cGG) // cardG_gt1.
+by apply/(lcn_nil_classP 1); rewrite ?abelian_nil //; apply/commG1P.
+Qed.
+
+Lemma cprod_nil A B G : A \* B = G -> nilpotent G = nilpotent A && nilpotent B.
+Proof.
+move=> defG; case/cprodP: defG (defG) => [[H K -> ->{A B}] defG _] defGc.
+apply/idP/andP=> [nilG | [/lcnP[m LmH1] /lcnP[n LnK1]]].
+ by rewrite !(nilpotentS _ nilG) // -defG (mulG_subr, mulG_subl).
+apply/lcnP; exists (m + n.+1); apply/trivgP.
+case/cprodP: (lcn_cprod (m.+1 + n.+1) defGc) => _ <- _.
+by rewrite mulG_subG /= -{1}LmH1 -LnK1 !lcn_sub_leq ?leq_addl ?leq_addr.
+Qed.
+
+Lemma mulg_nil G H :
+ H \subset 'C(G) -> nilpotent (G * H) = nilpotent G && nilpotent H.
+Proof. by move=> cGH; rewrite -(cprod_nil (cprodEY cGH)) /= cent_joinEr. Qed.
+
+Lemma dprod_nil A B G : A \x B = G -> nilpotent G = nilpotent A && nilpotent B.
+Proof. by case/dprodP=> [[H K -> ->] <- cHK _]; rewrite mulg_nil.
+Qed.
+
+Lemma bigdprod_nil I r (P : pred I) (A_ : I -> {set gT}) G :
+ \big[dprod/1]_(i <- r | P i) A_ i = G
+ -> (forall i, P i -> nilpotent (A_ i)) -> nilpotent G.
+Proof.
+move=> defG nilA; elim/big_rec: _ => [|i B Pi nilB] in G defG *.
+ by rewrite -defG nilpotent1.
+have [[_ H _ defB] _ _ _] := dprodP defG.
+by rewrite (dprod_nil defG) nilA //= defB nilB.
+Qed.
+
+End LowerCentral.
+
+Notation "''L_' n ( G )" := (lower_central_at_group n G) : Group_scope.
+
+Lemma lcn_cont n : GFunctor.continuous (lower_central_at n).
+Proof.
+case: n => //; elim=> // n IHn g0T h0T H phi.
+by rewrite !lcnSn morphimR ?lcn_sub // commSg ?IHn.
+Qed.
+
+Canonical lcn_igFun n := [igFun by lcn_sub^~ n & lcn_cont n].
+Canonical lcn_gFun n := [gFun by lcn_cont n].
+Canonical lcn_mgFun n := [mgFun by fun _ G H => @lcnS _ n G H].
+
+Section UpperCentralFunctor.
+
+Variable n : nat.
+Implicit Type gT : finGroupType.
+
+Lemma ucn_pmap : exists hZ : GFunctor.pmap, @upper_central_at n = hZ.
+Proof.
+elim: n => [|n' [hZ defZ]]; first by exists trivGfun_pgFun.
+by exists [pgFun of center %% hZ]; rewrite /= -defZ.
+Qed.
+
+(* Now extract all the intermediate facts of the last proof. *)
+
+Lemma ucn_group_set gT (G : {group gT}) : group_set 'Z_n(G).
+Proof. by have [hZ ->] := ucn_pmap; exact: groupP. Qed.
+
+Canonical upper_central_at_group gT G := Group (@ucn_group_set gT G).
+
+Lemma ucn_sub gT (G : {group gT}) : 'Z_n(G) \subset G.
+Proof. by have [hZ ->] := ucn_pmap; exact: gFsub. Qed.
+
+Lemma morphim_ucn : GFunctor.pcontinuous (upper_central_at n).
+Proof. by have [hZ ->] := ucn_pmap; exact: pmorphimF. Qed.
+
+Canonical ucn_igFun := [igFun by ucn_sub & morphim_ucn].
+Canonical ucn_gFun := [gFun by morphim_ucn].
+Canonical ucn_pgFun := [pgFun by morphim_ucn].
+
+Variable (gT : finGroupType) (G : {group gT}).
+
+Lemma ucn_char : 'Z_n(G) \char G. Proof. exact: gFchar. Qed.
+Lemma ucn_norm : G \subset 'N('Z_n(G)). Proof. exact: gFnorm. Qed.
+Lemma ucn_normal : 'Z_n(G) <| G. Proof. exact: gFnormal. Qed.
+
+End UpperCentralFunctor.
+
+Notation "''Z_' n ( G )" := (upper_central_at_group n G) : Group_scope.
+
+Section UpperCentral.
+
+Variable gT : finGroupType.
+Implicit Types (A B : {set gT}) (G H : {group gT}).
+
+Lemma ucn0 A : 'Z_0(A) = 1.
+Proof. by []. Qed.
+
+Lemma ucnSn n A : 'Z_n.+1(A) = coset 'Z_n(A) @*^-1 'Z(A / 'Z_n(A)).
+Proof. by []. Qed.
+
+Lemma ucnE n A : 'Z_n(A) = upper_central_at_rec n A.
+Proof. by []. Qed.
+
+Lemma ucn_subS n G : 'Z_n(G) \subset 'Z_n.+1(G).
+Proof. by rewrite -{1}['Z_n(G)]ker_coset morphpreS ?sub1G. Qed.
+
+Lemma ucn_sub_geq m n G : n >= m -> 'Z_m(G) \subset 'Z_n(G).
+Proof.
+move/subnK <-; elim: {n}(n - m) => // n IHn.
+exact: subset_trans (ucn_subS _ _).
+Qed.
+
+Lemma ucn_central n G : 'Z_n.+1(G) / 'Z_n(G) = 'Z(G / 'Z_n(G)).
+Proof. by rewrite ucnSn cosetpreK. Qed.
+
+Lemma ucn_normalS n G : 'Z_n(G) <| 'Z_n.+1(G).
+Proof. by rewrite (normalS _ _ (ucn_normal n G)) ?ucn_subS ?ucn_sub. Qed.
+
+Lemma ucn_comm n G : [~: 'Z_n.+1(G), G] \subset 'Z_n(G).
+Proof.
+rewrite -quotient_cents2 ?normal_norm ?ucn_normal ?ucn_normalS //.
+by rewrite ucn_central subsetIr.
+Qed.
+
+Lemma ucn1 G : 'Z_1(G) = 'Z(G).
+Proof.
+apply: (quotient_inj (normal1 _) (normal1 _)).
+by rewrite /= (ucn_central 0) -injmF ?norms1 ?coset1_injm.
+Qed.
+
+Lemma ucnSnR n G : 'Z_n.+1(G) = [set x in G | [~: [set x], G] \subset 'Z_n(G)].
+Proof.
+apply/setP=> x; rewrite inE -(setIidPr (ucn_sub n.+1 G)) inE ucnSn.
+case Gx: (x \in G) => //=; have nZG := ucn_norm n G.
+rewrite -sub1set -sub_quotient_pre -?quotient_cents2 ?sub1set ?(subsetP nZG) //.
+by rewrite subsetI quotientS ?sub1set.
+Qed.
+
+Lemma ucn_cprod n A B G : A \* B = G -> 'Z_n(A) \* 'Z_n(B) = 'Z_n(G).
+Proof.
+case/cprodP=> [[H K -> ->{A B}] mulHK cHK].
+elim: n => [|n /cprodP[_ /= defZ cZn]]; first exact: cprod1g.
+set Z := 'Z_n(G) in defZ cZn; rewrite (ucnSn n G) /= -/Z.
+have /mulGsubP[nZH nZK]: H * K \subset 'N(Z) by rewrite mulHK gFnorm.
+have <-: 'Z(H / Z) * 'Z(K / Z) = 'Z(G / Z).
+ by rewrite -mulHK quotientMl // center_prod ?quotient_cents.
+have ZquoZ (B A : {group gT}):
+ B \subset 'C(A) -> 'Z_n(A) * 'Z_n(B) = Z -> 'Z(A / Z) = 'Z_n.+1(A) / Z.
+- move=> cAB {defZ}defZ; have cAZnB := subset_trans (ucn_sub n B) cAB.
+ have /second_isom[/=]: A \subset 'N(Z).
+ by rewrite -defZ normsM ?gFnorm ?cents_norm // centsC.
+ suffices ->: Z :&: A = 'Z_n(A).
+ by move=> f inj_f im_f; rewrite -!im_f ?gFsub // ucn_central injm_center.
+ rewrite -defZ -group_modl ?gFsub //; apply/mulGidPl.
+ have [-> | n_gt0] := posnP n; first exact: subsetIl.
+ by apply: subset_trans (ucn_sub_geq A n_gt0); rewrite /= setIC ucn1 setIS.
+rewrite (ZquoZ H K) 1?centsC 1?(centC cZn) // {ZquoZ}(ZquoZ K H) //.
+have cZn1: 'Z_n.+1(K) \subset 'C('Z_n.+1(H)) by apply: centSS cHK; apply: gFsub.
+rewrite -quotientMl ?quotientK ?mul_subG ?(subset_trans (gFsub _ _)) //=.
+rewrite cprodE // -cent_joinEr ?mulSGid //= cent_joinEr //= -/Z.
+by rewrite -defZ mulgSS ?ucn_subS.
+Qed.
+
+Lemma ucn_dprod n A B G : A \x B = G -> 'Z_n(A) \x 'Z_n(B) = 'Z_n(G).
+Proof.
+move=> defG; have [[K H defA defB] _ _ tiAB] := dprodP defG.
+rewrite !dprodEcp // in defG *; first exact: ucn_cprod.
+by rewrite defA defB; apply/trivgP; rewrite -tiAB defA defB setISS ?ucn_sub.
+Qed.
+
+Lemma ucn_bigcprod n I r P (F : I -> {set gT}) G :
+ \big[cprod/1]_(i <- r | P i) F i = G ->
+ \big[cprod/1]_(i <- r | P i) 'Z_n(F i) = 'Z_n(G).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first by rewrite gF1.
+by rewrite -(ucn_cprod n dG); have [[_ H _ dH]] := cprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma ucn_bigdprod n I r P (F : I -> {set gT}) G :
+ \big[dprod/1]_(i <- r | P i) F i = G ->
+ \big[dprod/1]_(i <- r | P i) 'Z_n(F i) = 'Z_n(G).
+Proof.
+elim/big_rec2: _ G => [_ <- | i A Z _ IH G dG]; first by rewrite gF1.
+by rewrite -(ucn_dprod n dG); have [[_ H _ dH]] := dprodP dG; rewrite dH (IH H).
+Qed.
+
+Lemma ucn_lcnP n G : ('L_n.+1(G) == 1) = ('Z_n(G) == G).
+Proof.
+rewrite !eqEsubset sub1G ucn_sub /= andbT -(ucn0 G).
+elim: {1 3}n 0 (addn0 n) => [j <- //|i IHi j].
+rewrite addSnnS => /IHi <- {IHi}; rewrite ucnSn lcnSn.
+have nZG := normal_norm (ucn_normal j G).
+have nZL := subset_trans (lcn_sub _ _) nZG.
+by rewrite -sub_morphim_pre // subsetI morphimS ?lcn_sub // quotient_cents2.
+Qed.
+
+Lemma ucnP G : reflect (exists n, 'Z_n(G) = G) (nilpotent G).
+Proof.
+apply: (iffP (lcnP G)) => [] [n /eqP clGn];
+ by exists n; apply/eqP; rewrite ucn_lcnP in clGn *.
+Qed.
+
+Lemma ucn_nil_classP n G :
+ nilpotent G -> reflect ('Z_n(G) = G) (nil_class G <= n).
+Proof.
+move=> nilG; rewrite (sameP (lcn_nil_classP n nilG) eqP) ucn_lcnP; exact: eqP.
+Qed.
+
+Lemma ucn_id n G : 'Z_n('Z_n(G)) = 'Z_n(G).
+Proof. by rewrite -{2}['Z_n(G)]gFid. Qed.
+
+Lemma ucn_nilpotent n G : nilpotent 'Z_n(G).
+Proof. by apply/ucnP; exists n; rewrite ucn_id. Qed.
+
+Lemma nil_class_ucn n G : nil_class 'Z_n(G) <= n.
+Proof. by apply/ucn_nil_classP; rewrite ?ucn_nilpotent ?ucn_id. Qed.
+
+End UpperCentral.
+
+Section MorphNil.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Implicit Type G : {group aT}.
+
+Lemma morphim_lcn n G : G \subset D -> f @* 'L_n(G) = 'L_n(f @* G).
+Proof.
+move=> sHG; case: n => //; elim=> // n IHn.
+by rewrite !lcnSn -IHn morphimR // (subset_trans _ sHG) // lcn_sub.
+Qed.
+
+Lemma injm_ucn n G : 'injm f -> G \subset D -> f @* 'Z_n(G) = 'Z_n(f @* G).
+Proof. exact: injmF. Qed.
+
+Lemma morphim_nil G : nilpotent G -> nilpotent (f @* G).
+Proof.
+case/ucnP=> n ZnG; apply/ucnP; exists n; apply/eqP.
+by rewrite eqEsubset ucn_sub /= -{1}ZnG morphim_ucn.
+Qed.
+
+Lemma injm_nil G : 'injm f -> G \subset D -> nilpotent (f @* G) = nilpotent G.
+Proof.
+move=> injf sGD; apply/idP/idP; last exact: morphim_nil.
+case/ucnP=> n; rewrite -injm_ucn // => /injm_morphim_inj defZ.
+by apply/ucnP; exists n; rewrite defZ ?(subset_trans (ucn_sub n G)).
+Qed.
+
+Lemma nil_class_morphim G : nilpotent G -> nil_class (f @* G) <= nil_class G.
+Proof.
+move=> nilG; rewrite (sameP (ucn_nil_classP _ (morphim_nil nilG)) eqP) /=.
+by rewrite eqEsubset ucn_sub -{1}(ucn_nil_classP _ nilG (leqnn _)) morphim_ucn.
+Qed.
+
+Lemma nil_class_injm G :
+ 'injm f -> G \subset D -> nil_class (f @* G) = nil_class G.
+Proof.
+move=> injf sGD; case nilG: (nilpotent G).
+ apply/eqP; rewrite eqn_leq nil_class_morphim //.
+ rewrite (sameP (lcn_nil_classP _ nilG) eqP) -subG1.
+ rewrite -(injmSK injf) ?(subset_trans (lcn_sub _ _)) // morphim1.
+ by rewrite morphim_lcn // (lcn_nil_classP _ _ (leqnn _)) //= injm_nil.
+transitivity #|G|; apply/eqP; rewrite eqn_leq.
+ rewrite -(card_injm injf sGD) (leq_trans (index_size _ _)) ?size_mkseq //.
+ by rewrite leqNgt -nilpotent_class injm_nil ?nilG.
+rewrite (leq_trans (index_size _ _)) ?size_mkseq // leqNgt -nilpotent_class.
+by rewrite nilG.
+Qed.
+
+End MorphNil.
+
+Section QuotientNil.
+
+Variables gT : finGroupType.
+Implicit Types (rT : finGroupType) (G H : {group gT}).
+
+Lemma quotient_ucn_add m n G : 'Z_(m + n)(G) / 'Z_n(G) = 'Z_m(G / 'Z_n(G)).
+Proof.
+elim: m => [|m IHm]; first exact: trivg_quotient.
+apply/setP=> Zx; have [x Nx ->{Zx}] := cosetP Zx.
+have [sZG nZG] := andP (ucn_normal n G).
+rewrite (ucnSnR m) inE -!sub1set -morphim_set1 //= -quotientR ?sub1set // -IHm.
+rewrite !quotientSGK ?(ucn_sub_geq, leq_addl, comm_subG _ nZG, sub1set) //=.
+by rewrite addSn /= ucnSnR inE.
+Qed.
+
+Lemma isog_nil rT G (L : {group rT}) : G \isog L -> nilpotent G = nilpotent L.
+Proof. by case/isogP=> f injf <-; rewrite injm_nil. Qed.
+
+Lemma isog_nil_class rT G (L : {group rT}) :
+ G \isog L -> nil_class G = nil_class L.
+Proof. by case/isogP=> f injf <-; rewrite nil_class_injm. Qed.
+
+Lemma quotient_nil G H : nilpotent G -> nilpotent (G / H).
+Proof. exact: morphim_nil. Qed.
+
+Lemma quotient_center_nil G : nilpotent (G / 'Z(G)) = nilpotent G.
+Proof.
+rewrite -ucn1; apply/idP/idP; last exact: quotient_nil.
+case/ucnP=> c nilGq; apply/ucnP; exists c.+1; have nsZ1G := ucn_normal 1 G.
+apply: (quotient_inj _ nsZ1G); last by rewrite /= -(addn1 c) quotient_ucn_add.
+by rewrite (normalS _ _ nsZ1G) ?ucn_sub ?ucn_sub_geq.
+Qed.
+
+Lemma nil_class_quotient_center G :
+ nilpotent (G) -> nil_class (G / 'Z(G)) = (nil_class G).-1.
+Proof.
+move=> nilG; have nsZ1G := ucn_normal 1 G.
+apply/eqP; rewrite -ucn1 eqn_leq; apply/andP; split.
+ apply/ucn_nil_classP; rewrite ?quotient_nil //= -quotient_ucn_add ucn1.
+ by rewrite (ucn_nil_classP _ _ _) ?addn1 ?leqSpred.
+rewrite -subn1 leq_subLR addnC; apply/ucn_nil_classP => //=.
+apply: (quotient_inj _ nsZ1G) => /=.
+ by apply: normalS (ucn_sub _ _) nsZ1G; rewrite /= addnS ucn_sub_geq.
+by rewrite quotient_ucn_add; apply/ucn_nil_classP; rewrite //= quotient_nil.
+Qed.
+
+Lemma nilpotent_sub_norm G H :
+ nilpotent G -> H \subset G -> 'N_G(H) \subset H -> G :=: H.
+Proof.
+move=> nilG sHG sNH; apply/eqP; rewrite eqEsubset sHG andbT; apply/negP=> nsGH.
+have{nsGH} [i sZH []]: exists2 i, 'Z_i(G) \subset H & ~ 'Z_i.+1(G) \subset H.
+ case/ucnP: nilG => n ZnG; rewrite -{}ZnG in nsGH.
+ elim: n => [|i IHi] in nsGH *; first by rewrite sub1G in nsGH.
+ by case sZH: ('Z_i(G) \subset H); [exists i | apply: IHi; rewrite sZH].
+apply: subset_trans sNH; rewrite subsetI ucn_sub -commg_subr.
+by apply: subset_trans sZH; apply: subset_trans (ucn_comm i G); exact: commgS.
+Qed.
+
+Lemma nilpotent_proper_norm G H :
+ nilpotent G -> H \proper G -> H \proper 'N_G(H).
+Proof.
+move=> nilG; rewrite properEneq properE subsetI normG => /andP[neHG sHG].
+by rewrite sHG; apply: contra neHG; move/(nilpotent_sub_norm nilG)->.
+Qed.
+
+Lemma nilpotent_subnormal G H : nilpotent G -> H \subset G -> H <|<| G.
+Proof.
+move=> nilG; elim: {H}_.+1 {-2}H (ltnSn (#|G| - #|H|)) => // m IHm H.
+rewrite ltnS => leGHm sHG; have:= sHG; rewrite subEproper.
+case/predU1P=> [->|]; first exact: subnormal_refl.
+move/(nilpotent_proper_norm nilG); set K := 'N_G(H) => prHK.
+have snHK: H <|<| K by rewrite normal_subnormal ?normalSG.
+have sKG: K \subset G by rewrite subsetIl.
+apply: subnormal_trans snHK (IHm _ (leq_trans _ leGHm) sKG).
+by rewrite ltn_sub2l ?proper_card ?(proper_sub_trans prHK).
+Qed.
+
+Lemma TI_center_nil G H : nilpotent G -> H <| G -> H :&: 'Z(G) = 1 -> H :=: 1.
+Proof.
+move=> nilG /andP[sHG nHG] tiHZ.
+rewrite -{1}(setIidPl sHG); have{nilG} /ucnP[n <-] := nilG.
+elim: n => [|n IHn]; apply/trivgP; rewrite ?subsetIr // -tiHZ.
+rewrite [H :&: 'Z(G)]setIA subsetI setIS ?ucn_sub //= (sameP commG1P trivgP).
+rewrite -commg_subr commGC in nHG.
+rewrite -IHn subsetI (subset_trans _ nHG) ?commSg ?subsetIl //=.
+by rewrite (subset_trans _ (ucn_comm n G)) ?commSg ?subsetIr.
+Qed.
+
+Lemma meet_center_nil G H :
+ nilpotent G -> H <| G -> H :!=: 1 -> H :&: 'Z(G) != 1.
+Proof. by move=> nilG nsHG; apply: contraNneq => /TI_center_nil->. Qed.
+
+Lemma center_nil_eq1 G : nilpotent G -> ('Z(G) == 1) = (G :==: 1).
+Proof.
+move=> nilG; apply/eqP/eqP=> [Z1 | ->]; last exact: center1.
+by rewrite (TI_center_nil nilG) // (setIidPr (center_sub G)).
+Qed.
+
+Lemma cyclic_nilpotent_quo_der1_cyclic G :
+ nilpotent G -> cyclic (G / G^`(1)) -> cyclic G.
+Proof.
+move=> nG; rewrite (isog_cyclic (quotient1_isog G)).
+have [-> // | ntG' cGG'] := (eqVneq G^`(1) 1)%g.
+suffices: 'L_2(G) \subset G :&: 'L_3(G) by move/(eqfun_inP nG)=> <-.
+rewrite subsetI lcn_sub /= -quotient_cents2 ?lcn_norm //.
+apply: cyclic_factor_abelian (lcn_central 2 G) _.
+by rewrite (isog_cyclic (third_isog _ _ _)) ?lcn_normal // lcn_subS.
+Qed.
+
+End QuotientNil.
+
+Section Solvable.
+
+Variable gT : finGroupType.
+Implicit Types G H : {group gT}.
+
+Lemma nilpotent_sol G : nilpotent G -> solvable G.
+Proof.
+move=> nilG; apply/forall_inP=> H /subsetIP[sHG sHH'].
+by rewrite (forall_inP nilG) // subsetI sHG (subset_trans sHH') ?commgS.
+Qed.
+
+Lemma abelian_sol G : abelian G -> solvable G.
+Proof. by move/abelian_nil; exact: nilpotent_sol. Qed.
+
+Lemma solvable1 : solvable [1 gT]. Proof. exact: abelian_sol (abelian1 gT). Qed.
+
+Lemma solvableS G H : H \subset G -> solvable G -> solvable H.
+Proof.
+move=> sHG solG; apply/forall_inP=> K /subsetIP[sKH sKK'].
+by rewrite (forall_inP solG) // subsetI (subset_trans sKH).
+Qed.
+
+Lemma sol_der1_proper G H :
+ solvable G -> H \subset G -> H :!=: 1 -> H^`(1) \proper H.
+Proof.
+move=> solG sHG ntH; rewrite properE comm_subG //; apply: implyP ntH.
+by have:= forallP solG H; rewrite subsetI sHG implybNN.
+Qed.
+
+Lemma derivedP G : reflect (exists n, G^`(n) = 1) (solvable G).
+Proof.
+apply: (iffP idP) => [solG | [n solGn]]; last first.
+ apply/forall_inP=> H /subsetIP[sHG sHH'].
+ rewrite -subG1 -{}solGn; elim: n => // n IHn.
+ exact: subset_trans sHH' (commgSS _ _).
+suffices IHn n: #|G^`(n)| <= (#|G|.-1 - n).+1.
+ by exists #|G|.-1; rewrite [G^`(_)]card_le1_trivg ?(leq_trans (IHn _)) ?subnn.
+elim: n => [|n IHn]; first by rewrite subn0 prednK.
+rewrite dergSn subnS -ltnS.
+have [-> | ntGn] := eqVneq G^`(n) 1; first by rewrite commG1 cards1.
+case: (_ - _) IHn => [|n']; first by rewrite leqNgt cardG_gt1 ntGn.
+by apply: leq_trans (proper_card _); exact: sol_der1_proper (der_sub _ _) _.
+Qed.
+
+End Solvable.
+
+Section MorphSol.
+
+Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+Variable G : {group gT}.
+
+Lemma morphim_sol : solvable G -> solvable (f @* G).
+Proof.
+move/(solvableS (subsetIr D G)); case/derivedP=> n Gn1; apply/derivedP.
+by exists n; rewrite /= -morphimIdom -morphim_der ?subsetIl // Gn1 morphim1.
+Qed.
+
+Lemma injm_sol : 'injm f -> G \subset D -> solvable (f @* G) = solvable G.
+Proof.
+move=> injf sGD; apply/idP/idP; last exact: morphim_sol.
+case/derivedP=> n Gn1; apply/derivedP; exists n; apply/trivgP.
+rewrite -(injmSK injf) ?(subset_trans (der_sub _ _)) ?morphim_der //.
+by rewrite Gn1 morphim1.
+Qed.
+
+End MorphSol.
+
+Section QuotientSol.
+
+Variables gT rT : finGroupType.
+Implicit Types G H K : {group gT}.
+
+Lemma isog_sol G (L : {group rT}) : G \isog L -> solvable G = solvable L.
+Proof. by case/isogP=> f injf <-; rewrite injm_sol. Qed.
+
+Lemma quotient_sol G H : solvable G -> solvable (G / H).
+Proof. exact: morphim_sol. Qed.
+
+Lemma series_sol G H : H <| G -> solvable G = solvable H && solvable (G / H).
+Proof.
+case/andP=> sHG nHG; apply/idP/andP=> [solG | [solH solGH]].
+ by rewrite quotient_sol // (solvableS sHG).
+apply/forall_inP=> K /subsetIP[sKG sK'K].
+suffices sKH: K \subset H by rewrite (forall_inP solH) // subsetI sKH.
+have nHK := subset_trans sKG nHG.
+rewrite -quotient_sub1 // subG1 (forall_inP solGH) //.
+by rewrite subsetI -morphimR ?morphimS.
+Qed.
+
+Lemma metacyclic_sol G : metacyclic G -> solvable G.
+Proof.
+case/metacyclicP=> K [cycK nsKG cycGq].
+by rewrite (series_sol nsKG) !abelian_sol ?cyclic_abelian.
+Qed.
+
+End QuotientSol.
diff --git a/mathcomp/solvable/pgroup.v b/mathcomp/solvable/pgroup.v
new file mode 100644
index 0000000..c6db976
--- /dev/null
+++ b/mathcomp/solvable/pgroup.v
@@ -0,0 +1,1355 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div.
+Require Import fintype bigop finset prime fingroup morphism.
+Require Import gfunctor automorphism quotient action gproduct cyclic.
+
+(******************************************************************************)
+(* Standard group notions and constructions based on the prime decomposition *)
+(* of the order of the group or its elements: *)
+(* pi.-group G <=> G is a pi-group, i.e., pi.-nat #|G|. *)
+(* -> Recall that here and in the sequel pi can be a single prime p. *)
+(* pi.-subgroup(H) G <=> H is a pi-subgroup of G. *)
+(* := (H \subset G) && pi.-group H. *)
+(* -> This is provided mostly as a shorhand, with few associated lemmas. *)
+(* However, we do establish some results on maximal pi-subgroups. *)
+(* pi.-elt x <=> x is a pi-element. *)
+(* := pi.-nat #[x] or pi.-group <[x]>. *)
+(* x.`_pi == the pi-constituent of x: the (unique) pi-element *)
+(* y \in <[x]> such that x * y^-1 is a pi'-element. *)
+(* pi.-Hall(G) H <=> H is a Hall pi-subgroup of G. *)
+(* := [&& H \subset G, pi.-group H & pi^'.-nat #|G : H|]. *)
+(* -> This is also eqivalent to H \subset G /\ #|H| = #|G|`_pi. *)
+(* p.-Sylow(G) P <=> P is a Sylow p-subgroup of G. *)
+(* -> This is the display and preferred input notation for p.-Hall(G) P. *)
+(* 'Syl_p(G) == the set of the p-Sylow subgroups of G. *)
+(* := [set P : {group _} | p.-Sylow(G) P]. *)
+(* p_group P <=> P is a p-group for some prime p. *)
+(* Hall G H <=> H is a Hall pi-subgroup of G for some pi. *)
+(* := coprime #|H| #|G : H| && (H \subset G). *)
+(* Sylow G P <=> P is a Sylow p-subgroup of G for some p. *)
+(* := p_group P && Hall G P. *)
+(* 'O_pi(G) == the pi-core (largest normal pi-subgroup) of G. *)
+(* pcore_mod pi G H == the pi-core of G mod H. *)
+(* := G :&: (coset H @*^-1 'O_pi(G / H)). *)
+(* 'O_{pi2, pi1}(G) == the pi1,pi2-core of G. *)
+(* := the pi1-core of G mod 'O_pi2(G). *)
+(* -> We have 'O_{pi2, pi1}(G) / 'O_pi2(G) = 'O_pi1(G / 'O_pi2(G)) *)
+(* with 'O_pi2(G) <| 'O_{pi2, pi1}(G) <| G. *)
+(* 'O_{pn, ..., p1}(G) == the p1, ..., pn-core of G. *)
+(* := the p1-core of G mod 'O_{pn, ..., p2}(G). *)
+(* Note that notions are always defined on sets even though their name *)
+(* indicates "group" properties; the actual definition of the notion never *)
+(* tests for the group property, since this property will always be provided *)
+(* by a (canonical) group structure. Similarly, p-group properties assume *)
+(* without test that p is a prime. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section PgroupDefs.
+
+(* We defer the definition of the functors ('0_p(G), etc) because they need *)
+(* to quantify over the finGroupType explicitly. *)
+
+Variable gT : finGroupType.
+Implicit Type (x : gT) (A B : {set gT}) (pi : nat_pred) (p n : nat).
+
+Definition pgroup pi A := pi.-nat #|A|.
+
+Definition psubgroup pi A B := (B \subset A) && pgroup pi B.
+
+Definition p_group A := pgroup (pdiv #|A|) A.
+
+Definition p_elt pi x := pi.-nat #[x].
+
+Definition constt x pi := x ^+ (chinese #[x]`_pi #[x]`_pi^' 1 0).
+
+Definition Hall A B := (B \subset A) && coprime #|B| #|A : B|.
+
+Definition pHall pi A B := [&& B \subset A, pgroup pi B & pi^'.-nat #|A : B|].
+
+Definition Syl p A := [set P : {group gT} | pHall p A P].
+
+Definition Sylow A B := p_group B && Hall A B.
+
+End PgroupDefs.
+
+Arguments Scope pgroup [_ nat_scope group_scope].
+Arguments Scope psubgroup [_ nat_scope group_scope group_scope].
+Arguments Scope p_group [_ group_scope].
+Arguments Scope p_elt [_ nat_scope].
+Arguments Scope constt [_ group_scope nat_scope].
+Arguments Scope Hall [_ group_scope group_scope].
+Arguments Scope pHall [_ nat_scope group_scope group_scope].
+Arguments Scope Syl [_ nat_scope group_scope].
+Arguments Scope Sylow [_ group_scope group_scope].
+Prenex Implicits p_group Hall Sylow.
+
+Notation "pi .-group" := (pgroup pi)
+ (at level 2, format "pi .-group") : group_scope.
+
+Notation "pi .-subgroup ( A )" := (psubgroup pi A)
+ (at level 8, format "pi .-subgroup ( A )") : group_scope.
+
+Notation "pi .-elt" := (p_elt pi)
+ (at level 2, format "pi .-elt") : group_scope.
+
+Notation "x .`_ pi" := (constt x pi)
+ (at level 3, format "x .`_ pi") : group_scope.
+
+Notation "pi .-Hall ( G )" := (pHall pi G)
+ (at level 8, format "pi .-Hall ( G )") : group_scope.
+
+Notation "p .-Sylow ( G )" := (nat_pred_of_nat p).-Hall(G)
+ (at level 8, format "p .-Sylow ( G )") : group_scope.
+
+Notation "''Syl_' p ( G )" := (Syl p G)
+ (at level 8, p at level 2, format "''Syl_' p ( G )") : group_scope.
+
+Section PgroupProps.
+
+Variable gT : finGroupType.
+Implicit Types (pi rho : nat_pred) (p : nat).
+Implicit Types (x y z : gT) (A B C D : {set gT}) (G H K P Q R : {group gT}).
+
+Lemma trivgVpdiv G : G :=: 1 \/ (exists2 p, prime p & p %| #|G|).
+Proof.
+have [leG1|lt1G] := leqP #|G| 1; first by left; exact: card_le1_trivg.
+by right; exists (pdiv #|G|); rewrite ?pdiv_dvd ?pdiv_prime.
+Qed.
+
+Lemma prime_subgroupVti G H : prime #|G| -> G \subset H \/ H :&: G = 1.
+Proof.
+move=> prG; have [|[p p_pr pG]] := trivgVpdiv (H :&: G); first by right.
+left; rewrite (sameP setIidPr eqP) eqEcard subsetIr.
+suffices <-: p = #|G| by rewrite dvdn_leq ?cardG_gt0.
+by apply/eqP; rewrite -dvdn_prime2 // -(LagrangeI G H) setIC dvdn_mulr.
+Qed.
+
+Lemma pgroupE pi A : pi.-group A = pi.-nat #|A|. Proof. by []. Qed.
+
+Lemma sub_pgroup pi rho A : {subset pi <= rho} -> pi.-group A -> rho.-group A.
+Proof. by move=> pi_sub_rho; exact: sub_in_pnat (in1W pi_sub_rho). Qed.
+
+Lemma eq_pgroup pi rho A : pi =i rho -> pi.-group A = rho.-group A.
+Proof. exact: eq_pnat. Qed.
+
+Lemma eq_p'group pi rho A : pi =i rho -> pi^'.-group A = rho^'.-group A.
+Proof. by move/eq_negn; exact: eq_pnat. Qed.
+
+Lemma pgroupNK pi A : pi^'^'.-group A = pi.-group A.
+Proof. exact: pnatNK. Qed.
+
+Lemma pi_pgroup p pi A : p.-group A -> p \in pi -> pi.-group A.
+Proof. exact: pi_pnat. Qed.
+
+Lemma pi_p'group p pi A : pi.-group A -> p \in pi^' -> p^'.-group A.
+Proof. exact: pi_p'nat. Qed.
+
+Lemma pi'_p'group p pi A : pi^'.-group A -> p \in pi -> p^'.-group A.
+Proof. exact: pi'_p'nat. Qed.
+
+Lemma p'groupEpi p G : p^'.-group G = (p \notin \pi(G)).
+Proof. exact: p'natEpi (cardG_gt0 G). Qed.
+
+Lemma pgroup_pi G : \pi(G).-group G.
+Proof. by rewrite /=; exact: pnat_pi. Qed.
+
+Lemma partG_eq1 pi G : (#|G|`_pi == 1%N) = pi^'.-group G.
+Proof. exact: partn_eq1 (cardG_gt0 G). Qed.
+
+Lemma pgroupP pi G :
+ reflect (forall p, prime p -> p %| #|G| -> p \in pi) (pi.-group G).
+Proof. exact: pnatP. Qed.
+Implicit Arguments pgroupP [pi G].
+
+Lemma pgroup1 pi : pi.-group [1 gT].
+Proof. by rewrite /pgroup cards1. Qed.
+
+Lemma pgroupS pi G H : H \subset G -> pi.-group G -> pi.-group H.
+Proof. by move=> sHG; exact: pnat_dvd (cardSg sHG). Qed.
+
+Lemma oddSg G H : H \subset G -> odd #|G| -> odd #|H|.
+Proof. by rewrite !odd_2'nat; exact: pgroupS. Qed.
+
+Lemma odd_pgroup_odd p G : odd p -> p.-group G -> odd #|G|.
+Proof.
+move=> p_odd pG; rewrite odd_2'nat (pi_pnat pG) // !inE.
+by case: eqP p_odd => // ->.
+Qed.
+
+Lemma card_pgroup p G : p.-group G -> #|G| = (p ^ logn p #|G|)%N.
+Proof. by move=> pG; rewrite -p_part part_pnat_id. Qed.
+
+Lemma properG_ltn_log p G H :
+ p.-group G -> H \proper G -> logn p #|H| < logn p #|G|.
+Proof.
+move=> pG; rewrite properEneq eqEcard andbC ltnNge => /andP[sHG].
+rewrite sHG /= {1}(card_pgroup pG) {1}(card_pgroup (pgroupS sHG pG)).
+by apply: contra; case: p {pG} => [|p] leHG; rewrite ?logn0 // leq_pexp2l.
+Qed.
+
+Lemma pgroupM pi G H : pi.-group (G * H) = pi.-group G && pi.-group H.
+Proof.
+have GH_gt0: 0 < #|G :&: H| := cardG_gt0 _.
+rewrite /pgroup -(mulnK #|_| GH_gt0) -mul_cardG -(LagrangeI G H) -mulnA.
+by rewrite mulKn // -(LagrangeI H G) setIC !pnat_mul andbCA; case: (pnat _).
+Qed.
+
+Lemma pgroupJ pi G x : pi.-group (G :^ x) = pi.-group G.
+Proof. by rewrite /pgroup cardJg. Qed.
+
+Lemma pgroup_p p P : p.-group P -> p_group P.
+Proof.
+case: (leqP #|P| 1); first by move=> /card_le1_trivg-> _; exact: pgroup1.
+move/pdiv_prime=> pr_q pgP; have:= pgroupP pgP _ pr_q (pdiv_dvd _).
+by rewrite /p_group => /eqnP->.
+Qed.
+
+Lemma p_groupP P : p_group P -> exists2 p, prime p & p.-group P.
+Proof.
+case: (ltnP 1 #|P|); first by move/pdiv_prime; exists (pdiv #|P|).
+move/card_le1_trivg=> -> _; exists 2 => //; exact: pgroup1.
+Qed.
+
+Lemma pgroup_pdiv p G :
+ p.-group G -> G :!=: 1 ->
+ [/\ prime p, p %| #|G| & exists m, #|G| = p ^ m.+1]%N.
+Proof.
+move=> pG; rewrite trivg_card1; case/p_groupP: (pgroup_p pG) => q q_pr qG.
+move/implyP: (pgroupP pG q q_pr); case/p_natP: qG => // [[|m] ->] //.
+by rewrite dvdn_exp // => /eqnP <- _; split; rewrite ?dvdn_exp //; exists m.
+Qed.
+
+Lemma coprime_p'group p K R :
+ coprime #|K| #|R| -> p.-group R -> R :!=: 1 -> p^'.-group K.
+Proof.
+move=> coKR pR ntR; have [p_pr _ [e oK]] := pgroup_pdiv pR ntR.
+by rewrite oK coprime_sym coprime_pexpl // prime_coprime // -p'natE in coKR.
+Qed.
+
+Lemma card_Hall pi G H : pi.-Hall(G) H -> #|H| = #|G|`_pi.
+Proof.
+case/and3P=> sHG piH pi'H; rewrite -(Lagrange sHG).
+by rewrite partnM ?Lagrange // part_pnat_id ?part_p'nat ?muln1.
+Qed.
+
+Lemma pHall_sub pi A B : pi.-Hall(A) B -> B \subset A.
+Proof. by case/andP. Qed.
+
+Lemma pHall_pgroup pi A B : pi.-Hall(A) B -> pi.-group B.
+Proof. by case/and3P. Qed.
+
+Lemma pHallP pi G H : reflect (H \subset G /\ #|H| = #|G|`_pi) (pi.-Hall(G) H).
+Proof.
+apply: (iffP idP) => [piH | [sHG oH]].
+ split; [exact: pHall_sub piH | exact: card_Hall].
+rewrite /pHall sHG -divgS // /pgroup oH.
+by rewrite -{2}(@partnC pi #|G|) ?mulKn ?part_pnat.
+Qed.
+
+Lemma pHallE pi G H : pi.-Hall(G) H = (H \subset G) && (#|H| == #|G|`_pi).
+Proof. by apply/pHallP/andP=> [] [->] /eqP. Qed.
+
+Lemma coprime_mulpG_Hall pi G K R :
+ K * R = G -> pi.-group K -> pi^'.-group R ->
+ pi.-Hall(G) K /\ pi^'.-Hall(G) R.
+Proof.
+move=> defG piK pi'R; apply/andP.
+rewrite /pHall piK -!divgS /= -defG ?mulG_subl ?mulg_subr //= pnatNK.
+by rewrite coprime_cardMg ?(pnat_coprime piK) // mulKn ?mulnK //; exact/and3P.
+Qed.
+
+Lemma coprime_mulGp_Hall pi G K R :
+ K * R = G -> pi^'.-group K -> pi.-group R ->
+ pi^'.-Hall(G) K /\ pi.-Hall(G) R.
+Proof.
+move=> defG pi'K piR; apply/andP; rewrite andbC; apply/andP.
+by apply: coprime_mulpG_Hall => //; rewrite -(comm_group_setP _) defG ?groupP.
+Qed.
+
+Lemma eq_in_pHall pi rho G H :
+ {in \pi(G), pi =i rho} -> pi.-Hall(G) H = rho.-Hall(G) H.
+Proof.
+move=> eq_pi_rho; apply: andb_id2l => sHG.
+congr (_ && _); apply: eq_in_pnat => p piHp.
+ by apply: eq_pi_rho; exact: (piSg sHG).
+by congr (~~ _); apply: eq_pi_rho; apply: (pi_of_dvd (dvdn_indexg G H)).
+Qed.
+
+Lemma eq_pHall pi rho G H : pi =i rho -> pi.-Hall(G) H = rho.-Hall(G) H.
+Proof. by move=> eq_pi_rho; exact: eq_in_pHall (in1W eq_pi_rho). Qed.
+
+Lemma eq_p'Hall pi rho G H : pi =i rho -> pi^'.-Hall(G) H = rho^'.-Hall(G) H.
+Proof. by move=> eq_pi_rho; exact: eq_pHall (eq_negn _). Qed.
+
+Lemma pHallNK pi G H : pi^'^'.-Hall(G) H = pi.-Hall(G) H.
+Proof. exact: eq_pHall (negnK _). Qed.
+
+Lemma subHall_Hall pi rho G H K :
+ rho.-Hall(G) H -> {subset pi <= rho} -> pi.-Hall(H) K -> pi.-Hall(G) K.
+Proof.
+move=> hallH pi_sub_rho hallK.
+rewrite pHallE (subset_trans (pHall_sub hallK) (pHall_sub hallH)) /=.
+by rewrite (card_Hall hallK) (card_Hall hallH) partn_part.
+Qed.
+
+Lemma subHall_Sylow pi p G H P :
+ pi.-Hall(G) H -> p \in pi -> p.-Sylow(H) P -> p.-Sylow(G) P.
+Proof.
+move=> hallH pi_p sylP; have [sHG piH _] := and3P hallH.
+rewrite pHallE (subset_trans (pHall_sub sylP) sHG) /=.
+by rewrite (card_Hall sylP) (card_Hall hallH) partn_part // => q; move/eqnP->.
+Qed.
+
+Lemma pHall_Hall pi A B : pi.-Hall(A) B -> Hall A B.
+Proof. by case/and3P=> sBA piB pi'B; rewrite /Hall sBA (pnat_coprime piB). Qed.
+
+Lemma Hall_pi G H : Hall G H -> \pi(H).-Hall(G) H.
+Proof.
+by case/andP=> sHG coHG /=; rewrite /pHall sHG /pgroup pnat_pi -?coprime_pi'.
+Qed.
+
+Lemma HallP G H : Hall G H -> exists pi, pi.-Hall(G) H.
+Proof. by exists \pi(H); exact: Hall_pi. Qed.
+
+Lemma sdprod_Hall G K H : K ><| H = G -> Hall G K = Hall G H.
+Proof.
+case/sdprod_context=> /andP[sKG _] sHG defG _ tiKH.
+by rewrite /Hall sKG sHG -!divgS // -defG TI_cardMg // coprime_sym mulKn ?mulnK.
+Qed.
+
+Lemma coprime_sdprod_Hall_l G K H : K ><| H = G -> coprime #|K| #|H| = Hall G K.
+Proof.
+case/sdprod_context=> /andP[sKG _] _ defG _ tiKH.
+by rewrite /Hall sKG -divgS // -defG TI_cardMg ?mulKn.
+Qed.
+
+Lemma coprime_sdprod_Hall_r G K H : K ><| H = G -> coprime #|K| #|H| = Hall G H.
+Proof.
+by move=> defG; rewrite (coprime_sdprod_Hall_l defG) (sdprod_Hall defG).
+Qed.
+
+Lemma compl_pHall pi K H G :
+ pi.-Hall(G) K -> (H \in [complements to K in G]) = pi^'.-Hall(G) H.
+Proof.
+move=> hallK; apply/complP/idP=> [[tiKH mulKH] | hallH].
+ have [_] := andP hallK; rewrite /pHall pnatNK -{3}(invGid G) -mulKH mulG_subr.
+ by rewrite invMG !indexMg -indexgI andbC -indexgI setIC tiKH !indexg1.
+have [[sKG piK _] [sHG pi'H _]] := (and3P hallK, and3P hallH).
+have tiKH: K :&: H = 1 := coprime_TIg (pnat_coprime piK pi'H).
+split=> //; apply/eqP; rewrite eqEcard mul_subG //= TI_cardMg //.
+by rewrite (card_Hall hallK) (card_Hall hallH) partnC.
+Qed.
+
+Lemma compl_p'Hall pi K H G :
+ pi^'.-Hall(G) K -> (H \in [complements to K in G]) = pi.-Hall(G) H.
+Proof. by move/compl_pHall->; exact: eq_pHall (negnK pi). Qed.
+
+Lemma sdprod_normal_p'HallP pi K H G :
+ K <| G -> pi^'.-Hall(G) H -> reflect (K ><| H = G) (pi.-Hall(G) K).
+Proof.
+move=> nsKG hallH; rewrite -(compl_p'Hall K hallH).
+exact: sdprod_normal_complP.
+Qed.
+
+Lemma sdprod_normal_pHallP pi K H G :
+ K <| G -> pi.-Hall(G) H -> reflect (K ><| H = G) (pi^'.-Hall(G) K).
+Proof.
+by move=> nsKG hallH; apply: sdprod_normal_p'HallP; rewrite ?pHallNK.
+Qed.
+
+Lemma pHallJ2 pi G H x : pi.-Hall(G :^ x) (H :^ x) = pi.-Hall(G) H.
+Proof. by rewrite !pHallE conjSg !cardJg. Qed.
+
+Lemma pHallJnorm pi G H x : x \in 'N(G) -> pi.-Hall(G) (H :^ x) = pi.-Hall(G) H.
+Proof. by move=> Nx; rewrite -{1}(normP Nx) pHallJ2. Qed.
+
+Lemma pHallJ pi G H x : x \in G -> pi.-Hall(G) (H :^ x) = pi.-Hall(G) H.
+Proof. by move=> Gx; rewrite -{1}(conjGid Gx) pHallJ2. Qed.
+
+Lemma HallJ G H x : x \in G -> Hall G (H :^ x) = Hall G H.
+Proof.
+by move=> Gx; rewrite /Hall -!divgI -{1 3}(conjGid Gx) conjSg -conjIg !cardJg.
+Qed.
+
+Lemma psubgroupJ pi G H x :
+ x \in G -> pi.-subgroup(G) (H :^ x) = pi.-subgroup(G) H.
+Proof. by move=> Gx; rewrite /psubgroup pgroupJ -{1}(conjGid Gx) conjSg. Qed.
+
+Lemma p_groupJ P x : p_group (P :^ x) = p_group P.
+Proof. by rewrite /p_group cardJg pgroupJ. Qed.
+
+Lemma SylowJ G P x : x \in G -> Sylow G (P :^ x) = Sylow G P.
+Proof. by move=> Gx; rewrite /Sylow p_groupJ HallJ. Qed.
+
+Lemma p_Sylow p G P : p.-Sylow(G) P -> Sylow G P.
+Proof.
+by move=> pP; rewrite /Sylow (pgroup_p (pHall_pgroup pP)) (pHall_Hall pP).
+Qed.
+
+Lemma pHall_subl pi G K H :
+ H \subset K -> K \subset G -> pi.-Hall(G) H -> pi.-Hall(K) H.
+Proof.
+move=> sHK sKG; rewrite /pHall sHK; case/and3P=> _ ->.
+by apply: pnat_dvd; exact: indexSg.
+Qed.
+
+Lemma Hall1 G : Hall G 1.
+Proof. by rewrite /Hall sub1G cards1 coprime1n. Qed.
+
+Lemma p_group1 : @p_group gT 1.
+Proof. by rewrite (@pgroup_p 2) ?pgroup1. Qed.
+
+Lemma Sylow1 G : Sylow G 1.
+Proof. by rewrite /Sylow p_group1 Hall1. Qed.
+
+Lemma SylowP G P : reflect (exists2 p, prime p & p.-Sylow(G) P) (Sylow G P).
+Proof.
+apply: (iffP idP) => [| [p _]]; last exact: p_Sylow.
+case/andP=> /p_groupP[p p_pr] /p_natP[[P1 _ | n oP /Hall_pi]]; last first.
+ by rewrite /= oP pi_of_exp // (eq_pHall _ _ (pi_of_prime _)) //; exists p.
+have{p p_pr P1} ->: P :=: 1 by apply: card1_trivg; rewrite P1.
+pose p := pdiv #|G|.+1; have p_pr: prime p by rewrite pdiv_prime ?ltnS.
+exists p; rewrite // pHallE sub1G cards1 part_p'nat //.
+apply/pgroupP=> q pr_q qG; apply/eqnP=> def_q.
+have: p %| #|G| + 1 by rewrite addn1 pdiv_dvd.
+by rewrite dvdn_addr -def_q // Euclid_dvd1.
+Qed.
+
+Lemma p_elt_exp pi x m : pi.-elt (x ^+ m) = (#[x]`_pi^' %| m).
+Proof.
+apply/idP/idP=> [pi_xm | /dvdnP[q ->{m}]]; last first.
+ rewrite mulnC; apply: pnat_dvd (part_pnat pi #[x]).
+ by rewrite order_dvdn -expgM mulnC mulnA partnC // -order_dvdn dvdn_mulr.
+rewrite -(@Gauss_dvdr _ #[x ^+ m]); last first.
+ by rewrite coprime_sym (pnat_coprime pi_xm) ?part_pnat.
+apply: (@dvdn_trans #[x]); first by rewrite -{2}[#[x]](partnC pi) ?dvdn_mull.
+by rewrite order_dvdn mulnC expgM expg_order.
+Qed.
+
+Lemma mem_p_elt pi x G : pi.-group G -> x \in G -> pi.-elt x.
+Proof. by move=> piG Gx; apply: pgroupS piG; rewrite cycle_subG. Qed.
+
+Lemma p_eltM_norm pi x y :
+ x \in 'N(<[y]>) -> pi.-elt x -> pi.-elt y -> pi.-elt (x * y).
+Proof.
+move=> nyx pi_x pi_y; apply: (@mem_p_elt pi _ (<[x]> <*> <[y]>)%G).
+ by rewrite /= norm_joinEl ?cycle_subG // pgroupM; exact/andP.
+by rewrite groupM // mem_gen // inE cycle_id ?orbT.
+Qed.
+
+Lemma p_eltM pi x y : commute x y -> pi.-elt x -> pi.-elt y -> pi.-elt (x * y).
+Proof.
+move=> cxy; apply: p_eltM_norm; apply: (subsetP (cent_sub _)).
+by rewrite cent_gen cent_set1; exact/cent1P.
+Qed.
+
+Lemma p_elt1 pi : pi.-elt (1 : gT).
+Proof. by rewrite /p_elt order1. Qed.
+
+Lemma p_eltV pi x : pi.-elt x^-1 = pi.-elt x.
+Proof. by rewrite /p_elt orderV. Qed.
+
+Lemma p_eltX pi x n : pi.-elt x -> pi.-elt (x ^+ n).
+Proof. by rewrite -{1}[x]expg1 !p_elt_exp dvdn1 => /eqnP->. Qed.
+
+Lemma p_eltJ pi x y : pi.-elt (x ^ y) = pi.-elt x.
+Proof. by congr pnat; rewrite orderJ. Qed.
+
+Lemma sub_p_elt pi1 pi2 x : {subset pi1 <= pi2} -> pi1.-elt x -> pi2.-elt x.
+Proof. by move=> pi12; apply: sub_in_pnat => q _; exact: pi12. Qed.
+
+Lemma eq_p_elt pi1 pi2 x : pi1 =i pi2 -> pi1.-elt x = pi2.-elt x.
+Proof. by move=> pi12; exact: eq_pnat. Qed.
+
+Lemma p_eltNK pi x : pi^'^'.-elt x = pi.-elt x.
+Proof. exact: pnatNK. Qed.
+
+Lemma eq_constt pi1 pi2 x : pi1 =i pi2 -> x.`_pi1 = x.`_pi2.
+Proof.
+move=> pi12; congr (x ^+ (chinese _ _ 1 0)); apply: eq_partn => // a.
+by congr (~~ _); exact: pi12.
+Qed.
+
+Lemma consttNK pi x : x.`_pi^'^' = x.`_pi.
+Proof. by rewrite /constt !partnNK. Qed.
+
+Lemma cycle_constt pi x : x.`_pi \in <[x]>.
+Proof. exact: mem_cycle. Qed.
+
+Lemma consttV pi x : (x^-1).`_pi = (x.`_pi)^-1.
+Proof. by rewrite /constt expgVn orderV. Qed.
+
+Lemma constt1 pi : 1.`_pi = 1 :> gT.
+Proof. exact: expg1n. Qed.
+
+Lemma consttJ pi x y : (x ^ y).`_pi = x.`_pi ^ y.
+Proof. by rewrite /constt orderJ conjXg. Qed.
+
+Lemma p_elt_constt pi x : pi.-elt x.`_pi.
+Proof. by rewrite p_elt_exp /chinese addn0 mul1n dvdn_mulr. Qed.
+
+Lemma consttC pi x : x.`_pi * x.`_pi^' = x.
+Proof.
+apply/eqP; rewrite -{3}[x]expg1 -expgD eq_expg_mod_order.
+rewrite partnNK -{5 6}(@partnC pi #[x]) // /chinese !addn0.
+by rewrite chinese_remainder ?chinese_modl ?chinese_modr ?coprime_partC ?eqxx.
+Qed.
+
+Lemma p'_elt_constt pi x : pi^'.-elt (x * (x.`_pi)^-1).
+Proof. by rewrite -{1}(consttC pi^' x) consttNK mulgK p_elt_constt. Qed.
+
+Lemma order_constt pi (x : gT) : #[x.`_pi] = #[x]`_pi.
+Proof.
+rewrite -{2}(consttC pi x) orderM; [|exact: commuteX2|]; last first.
+ by apply: (@pnat_coprime pi); exact: p_elt_constt.
+by rewrite partnM // part_pnat_id ?part_p'nat ?muln1 //; exact: p_elt_constt.
+Qed.
+
+Lemma consttM pi x y : commute x y -> (x * y).`_pi = x.`_pi * y.`_pi.
+Proof.
+move=> cxy; pose m := #|<<[set x; y]>>|; have m_gt0: 0 < m := cardG_gt0 _.
+pose k := chinese m`_pi m`_pi^' 1 0.
+suffices kXpi z: z \in <<[set x; y]>> -> z.`_pi = z ^+ k.
+ by rewrite !kXpi ?expgMn // ?groupM ?mem_gen // !inE eqxx ?orbT.
+move=> xyz; have{xyz} zm: #[z] %| m by rewrite cardSg ?cycle_subG.
+apply/eqP; rewrite eq_expg_mod_order -{3 4}[#[z]](partnC pi) //.
+rewrite chinese_remainder ?chinese_modl ?chinese_modr ?coprime_partC //.
+rewrite -!(modn_dvdm k (partn_dvd _ m_gt0 zm)).
+rewrite chinese_modl ?chinese_modr ?coprime_partC //.
+by rewrite !modn_dvdm ?partn_dvd ?eqxx.
+Qed.
+
+Lemma consttX pi x n : (x ^+ n).`_pi = x.`_pi ^+ n.
+Proof.
+elim: n => [|n IHn]; first exact: constt1.
+rewrite !expgS consttM ?IHn //; exact: commuteX.
+Qed.
+
+Lemma constt1P pi x : reflect (x.`_pi = 1) (pi^'.-elt x).
+Proof.
+rewrite -{2}[x]expg1 p_elt_exp -order_constt consttNK order_dvdn expg1.
+exact: eqP.
+Qed.
+
+Lemma constt_p_elt pi x : pi.-elt x -> x.`_pi = x.
+Proof.
+by rewrite -p_eltNK -{3}(consttC pi x) => /constt1P->; rewrite mulg1.
+Qed.
+
+Lemma sub_in_constt pi1 pi2 x :
+ {in \pi(#[x]), {subset pi1 <= pi2}} -> x.`_pi2.`_pi1 = x.`_pi1.
+Proof.
+move=> pi12; rewrite -{2}(consttC pi2 x) consttM; last exact: commuteX2.
+rewrite (constt1P _ x.`_pi2^' _) ?mulg1 //.
+apply: sub_in_pnat (p_elt_constt _ x) => p; rewrite order_constt => pi_p.
+apply: contra; apply: pi12.
+by rewrite -[#[x]](partnC pi2^') // primes_mul // pi_p.
+Qed.
+
+Lemma prod_constt x : \prod_(0 <= p < #[x].+1) x.`_p = x.
+Proof.
+pose lp n := [pred p | p < n].
+have: (lp #[x].+1).-elt x by apply/pnatP=> // p _; exact: dvdn_leq.
+move/constt_p_elt=> def_x; symmetry; rewrite -{1}def_x {def_x}.
+elim: _.+1 => [|p IHp].
+ by rewrite big_nil; apply/constt1P; exact/pgroupP.
+rewrite big_nat_recr //= -{}IHp -(consttC (lp p) x.`__); congr (_ * _).
+ rewrite sub_in_constt // => q _; exact: leqW.
+set y := _.`__; rewrite -(consttC p y) (constt1P p^' _ _) ?mulg1.
+ by rewrite 2?sub_in_constt // => q _; move/eqnP->; rewrite !inE ?ltnn.
+rewrite /p_elt pnatNK !order_constt -partnI.
+apply: sub_in_pnat (part_pnat _ _) => q _.
+by rewrite !inE ltnS -leqNgt -eqn_leq.
+Qed.
+
+Lemma max_pgroupJ pi M G x :
+ x \in G -> [max M | pi.-subgroup(G) M] ->
+ [max M :^ x of M | pi.-subgroup(G) M].
+Proof.
+move=> Gx /maxgroupP[piM maxM]; apply/maxgroupP.
+split=> [|H piH]; first by rewrite psubgroupJ.
+by rewrite -(conjsgKV x H) conjSg => /maxM/=-> //; rewrite psubgroupJ ?groupV.
+Qed.
+
+Lemma comm_sub_max_pgroup pi H M G :
+ [max M | pi.-subgroup(G) M] -> pi.-group H -> H \subset G ->
+ commute H M -> H \subset M.
+Proof.
+case/maxgroupP=> /andP[sMG piM] maxM piH sHG cHM.
+rewrite -(maxM (H <*> M)%G) /= comm_joingE ?(mulG_subl, mulG_subr) //.
+by rewrite /psubgroup pgroupM piM piH mul_subG.
+Qed.
+
+Lemma normal_sub_max_pgroup pi H M G :
+ [max M | pi.-subgroup(G) M] -> pi.-group H -> H <| G -> H \subset M.
+Proof.
+move=> maxM piH /andP[sHG nHG].
+apply: comm_sub_max_pgroup piH sHG _ => //; apply: commute_sym; apply: normC.
+by apply: subset_trans nHG; case/andP: (maxgroupp maxM).
+Qed.
+
+Lemma norm_sub_max_pgroup pi H M G :
+ [max M | pi.-subgroup(G) M] -> pi.-group H -> H \subset G ->
+ H \subset 'N(M) -> H \subset M.
+Proof. by move=> maxM piH sHG /normC; exact: comm_sub_max_pgroup piH sHG. Qed.
+
+Lemma sub_pHall pi H G K :
+ pi.-Hall(G) H -> pi.-group K -> H \subset K -> K \subset G -> K :=: H.
+Proof.
+move=> hallH piK sHK sKG; apply/eqP; rewrite eq_sym eqEcard sHK.
+by rewrite (card_Hall hallH) -(part_pnat_id piK) dvdn_leq ?partn_dvd ?cardSg.
+Qed.
+
+Lemma Hall_max pi H G : pi.-Hall(G) H -> [max H | pi.-subgroup(G) H].
+Proof.
+move=> hallH; apply/maxgroupP; split=> [|K].
+ by rewrite /psubgroup; case/and3P: hallH => ->.
+case/andP=> sKG piK sHK; exact: (sub_pHall hallH).
+Qed.
+
+Lemma pHall_id pi H G : pi.-Hall(G) H -> pi.-group G -> H :=: G.
+Proof.
+by move=> hallH piG; rewrite (sub_pHall hallH piG) ?(pHall_sub hallH).
+Qed.
+
+Lemma psubgroup1 pi G : pi.-subgroup(G) 1.
+Proof. by rewrite /psubgroup sub1G pgroup1. Qed.
+
+Lemma Cauchy p G : prime p -> p %| #|G| -> {x | x \in G & #[x] = p}.
+Proof.
+move=> p_pr; elim: {G}_.+1 {-2}G (ltnSn #|G|) => // n IHn G.
+rewrite ltnS => leGn pG; pose xpG := [pred x in G | #[x] == p].
+have [x /andP[Gx /eqP] | no_x] := pickP xpG; first by exists x.
+have{pG n leGn IHn} pZ: p %| #|'C_G(G)|.
+ suffices /dvdn_addl <-: p %| #|G :\: 'C(G)| by rewrite cardsID.
+ have /acts_sum_card_orbit <-: [acts G, on G :\: 'C(G) | 'J].
+ by apply/actsP=> x Gx y; rewrite !inE -!mem_conjgV -centJ conjGid ?groupV.
+ elim/big_rec: _ => // _ _ /imsetP[x /setDP[Gx nCx] ->] /dvdn_addl->.
+ have ltCx: 'C_G[x] \proper G by rewrite properE subsetIl subsetIidl sub_cent1.
+ have /negP: ~ p %| #|'C_G[x]|.
+ case/(IHn _ (leq_trans (proper_card ltCx) leGn))=> y /setIP[Gy _] /eqP-oy.
+ by have /andP[] := no_x y.
+ by apply/implyP; rewrite -index_cent1 indexgI implyNb -Euclid_dvdM ?LagrangeI.
+have [Q maxQ _]: {Q | [max Q | p^'.-subgroup('C_G(G)) Q] & 1%G \subset Q}.
+ apply: maxgroup_exists; exact: psubgroup1.
+case/andP: (maxgroupp maxQ) => sQC; rewrite /pgroup p'natE // => /negP[].
+apply: dvdn_trans pZ (cardSg _); apply/subsetP=> x /setIP[Gx Cx].
+rewrite -sub1set -gen_subG (normal_sub_max_pgroup maxQ) //; last first.
+ rewrite /normal subsetI !cycle_subG ?Gx ?cents_norm ?subIset ?andbT //=.
+ by rewrite centsC cycle_subG Cx.
+rewrite /pgroup p'natE //= -[#|_|]/#[x]; apply/dvdnP=> [[m oxm]].
+have m_gt0: 0 < m by apply: dvdn_gt0 (order_gt0 x) _; rewrite oxm dvdn_mulr.
+case/idP: (no_x (x ^+ m)); rewrite /= groupX //= orderXgcd //= oxm.
+by rewrite gcdnC gcdnMr mulKn.
+Qed.
+
+(* These lemmas actually hold for maximal pi-groups, but below we'll *)
+(* derive from the Cauchy lemma that a normal max pi-group is Hall. *)
+
+Lemma sub_normal_Hall pi G H K :
+ pi.-Hall(G) H -> H <| G -> K \subset G -> (K \subset H) = pi.-group K.
+Proof.
+move=> hallH nsHG sKG; apply/idP/idP=> [sKH | piK].
+ by rewrite (pgroupS sKH) ?(pHall_pgroup hallH).
+apply: norm_sub_max_pgroup (Hall_max hallH) piK _ _ => //.
+exact: subset_trans sKG (normal_norm nsHG).
+Qed.
+
+Lemma mem_normal_Hall pi H G x :
+ pi.-Hall(G) H -> H <| G -> x \in G -> (x \in H) = pi.-elt x.
+Proof. by rewrite -!cycle_subG; exact: sub_normal_Hall. Qed.
+
+Lemma uniq_normal_Hall pi H G K :
+ pi.-Hall(G) H -> H <| G -> [max K | pi.-subgroup(G) K] -> K :=: H.
+Proof.
+move=> hallH nHG /maxgroupP[/andP[sKG piK] /(_ H) -> //].
+ exact: (maxgroupp (Hall_max hallH)).
+by rewrite (sub_normal_Hall hallH).
+Qed.
+
+End PgroupProps.
+
+Implicit Arguments pgroupP [gT pi G].
+Implicit Arguments constt1P [gT pi x].
+Prenex Implicits pgroupP constt1P.
+
+Section NormalHall.
+
+Variables (gT : finGroupType) (pi : nat_pred).
+Implicit Types G H K : {group gT}.
+
+Lemma normal_max_pgroup_Hall G H :
+ [max H | pi.-subgroup(G) H] -> H <| G -> pi.-Hall(G) H.
+Proof.
+case/maxgroupP=> /andP[sHG piH] maxH nsHG; have [_ nHG] := andP nsHG.
+rewrite /pHall sHG piH; apply/pnatP=> // p p_pr.
+rewrite inE /= -pnatE // -card_quotient //.
+case/Cauchy=> //= Hx; rewrite -sub1set -gen_subG -/<[Hx]> /order.
+case/inv_quotientS=> //= K -> sHK sKG {Hx}.
+rewrite card_quotient ?(subset_trans sKG) // => iKH; apply/negP=> pi_p.
+rewrite -iKH -divgS // (maxH K) ?divnn ?cardG_gt0 // in p_pr.
+by rewrite /psubgroup sKG /pgroup -(Lagrange sHK) mulnC pnat_mul iKH pi_p.
+Qed.
+
+Lemma setI_normal_Hall G H K :
+ H <| G -> pi.-Hall(G) H -> K \subset G -> pi.-Hall(K) (H :&: K).
+Proof.
+move=> nsHG hallH sKG; apply: normal_max_pgroup_Hall; last first.
+ by rewrite /= setIC (normalGI sKG nsHG).
+apply/maxgroupP; split=> [|M /andP[sMK piM] sHK_M].
+ by rewrite /psubgroup subsetIr (pgroupS (subsetIl _ _) (pHall_pgroup hallH)).
+apply/eqP; rewrite eqEsubset sHK_M subsetI sMK !andbT.
+by rewrite (sub_normal_Hall hallH) // (subset_trans sMK).
+Qed.
+
+End NormalHall.
+
+Section Morphim.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Implicit Types (pi : nat_pred) (G H P : {group aT}).
+
+Lemma morphim_pgroup pi G : pi.-group G -> pi.-group (f @* G).
+Proof. by apply: pnat_dvd; exact: dvdn_morphim. Qed.
+
+Lemma morphim_odd G : odd #|G| -> odd #|f @* G|.
+Proof. by rewrite !odd_2'nat; exact: morphim_pgroup. Qed.
+
+Lemma pmorphim_pgroup pi G :
+ pi.-group ('ker f) -> G \subset D -> pi.-group (f @* G) = pi.-group G.
+Proof.
+move=> piker sGD; apply/idP/idP=> [pifG|]; last exact: morphim_pgroup.
+apply: (@pgroupS _ _ (f @*^-1 (f @* G))); first by rewrite -sub_morphim_pre.
+by rewrite /pgroup card_morphpre ?morphimS // pnat_mul; exact/andP.
+Qed.
+
+Lemma morphim_p_index pi G H :
+ H \subset D -> pi.-nat #|G : H| -> pi.-nat #|f @* G : f @* H|.
+Proof.
+by move=> sHD; apply: pnat_dvd; rewrite index_morphim ?subIset // sHD orbT.
+Qed.
+
+Lemma morphim_pHall pi G H :
+ H \subset D -> pi.-Hall(G) H -> pi.-Hall(f @* G) (f @* H).
+Proof.
+move=> sHD /and3P[sHG piH pi'GH].
+by rewrite /pHall morphimS // morphim_pgroup // morphim_p_index.
+Qed.
+
+Lemma pmorphim_pHall pi G H :
+ G \subset D -> H \subset D -> pi.-subgroup(H :&: G) ('ker f) ->
+ pi.-Hall(f @* G) (f @* H) = pi.-Hall(G) H.
+Proof.
+move=> sGD sHD /andP[/subsetIP[sKH sKG] piK]; rewrite !pHallE morphimSGK //.
+apply: andb_id2l => sHG; rewrite -(Lagrange sKH) -(Lagrange sKG) partnM //.
+by rewrite (part_pnat_id piK) !card_morphim !(setIidPr _) // eqn_pmul2l.
+Qed.
+
+Lemma morphim_Hall G H : H \subset D -> Hall G H -> Hall (f @* G) (f @* H).
+Proof.
+by move=> sHD /HallP[pi piH]; apply: (@pHall_Hall _ pi); exact: morphim_pHall.
+Qed.
+
+Lemma morphim_pSylow p G P :
+ P \subset D -> p.-Sylow(G) P -> p.-Sylow(f @* G) (f @* P).
+Proof. exact: morphim_pHall. Qed.
+
+Lemma morphim_p_group P : p_group P -> p_group (f @* P).
+Proof. by move/morphim_pgroup; exact: pgroup_p. Qed.
+
+Lemma morphim_Sylow G P : P \subset D -> Sylow G P -> Sylow (f @* G) (f @* P).
+Proof.
+by move=> sPD /andP[pP hallP]; rewrite /Sylow morphim_p_group // morphim_Hall.
+Qed.
+
+Lemma morph_p_elt pi x : x \in D -> pi.-elt x -> pi.-elt (f x).
+Proof. by move=> Dx; apply: pnat_dvd; exact: morph_order. Qed.
+
+Lemma morph_constt pi x : x \in D -> f x.`_pi = (f x).`_pi.
+Proof.
+move=> Dx; rewrite -{2}(consttC pi x) morphM ?groupX //.
+rewrite consttM; last by rewrite !morphX //; exact: commuteX2.
+have: pi.-elt (f x.`_pi) by rewrite morph_p_elt ?groupX ?p_elt_constt //.
+have: pi^'.-elt (f x.`_pi^') by rewrite morph_p_elt ?groupX ?p_elt_constt //.
+by move/constt1P->; move/constt_p_elt->; rewrite mulg1.
+Qed.
+
+End Morphim.
+
+Section Pquotient.
+
+Variables (pi : nat_pred) (gT : finGroupType) (p : nat) (G H K : {group gT}).
+Hypothesis piK : pi.-group K.
+
+Lemma quotient_pgroup : pi.-group (K / H). Proof. exact: morphim_pgroup. Qed.
+
+Lemma quotient_pHall :
+ K \subset 'N(H) -> pi.-Hall(G) K -> pi.-Hall(G / H) (K / H).
+Proof. exact: morphim_pHall. Qed.
+
+Lemma quotient_odd : odd #|K| -> odd #|K / H|. Proof. exact: morphim_odd. Qed.
+
+Lemma pquotient_pgroup : G \subset 'N(K) -> pi.-group (G / K) = pi.-group G.
+Proof. by move=> nKG; rewrite pmorphim_pgroup ?ker_coset. Qed.
+
+Lemma pquotient_pHall :
+ K <| G -> K <| H -> pi.-Hall(G / K) (H / K) = pi.-Hall(G) H.
+Proof.
+case/andP=> sKG nKG; case/andP=> sKH nKH.
+by rewrite pmorphim_pHall // ker_coset /psubgroup subsetI sKH sKG.
+Qed.
+
+Lemma ltn_log_quotient :
+ p.-group G -> H :!=: 1 -> H \subset G -> logn p #|G / H| < logn p #|G|.
+Proof.
+move=> pG ntH sHG; apply: contraLR (ltn_quotient ntH sHG); rewrite -!leqNgt.
+rewrite {2}(card_pgroup pG) {2}(card_pgroup (morphim_pgroup _ pG)).
+by case: (posnP p) => [-> //|]; exact: leq_pexp2l.
+Qed.
+
+End Pquotient.
+
+(* Application of card_Aut_cyclic to internal faithful action on cyclic *)
+(* p-subgroups. *)
+Section InnerAutCyclicPgroup.
+
+Variables (gT : finGroupType) (p : nat) (G C : {group gT}).
+Hypothesis nCG : G \subset 'N(C).
+
+Lemma logn_quotient_cent_cyclic_pgroup :
+ p.-group C -> cyclic C -> logn p #|G / 'C_G(C)| <= (logn p #|C|).-1.
+Proof.
+move=> pC cycC; have [-> | ntC] := eqsVneq C 1.
+ by rewrite cent1T setIT trivg_quotient cards1 logn1.
+have [p_pr _ [e oC]] := pgroup_pdiv pC ntC.
+rewrite -ker_conj_aut (card_isog (first_isog_loc _ _)) //.
+apply: leq_trans (dvdn_leq_log _ _ (cardSg (Aut_conj_aut _ _))) _ => //.
+rewrite card_Aut_cyclic // oC totient_pfactor //= logn_Gauss ?pfactorK //.
+by rewrite prime_coprime // gtnNdvd // -(subnKC (prime_gt1 p_pr)).
+Qed.
+
+Lemma p'group_quotient_cent_prime :
+ prime p -> #|C| %| p -> p^'.-group (G / 'C_G(C)).
+Proof.
+move=> p_pr pC; have pgC: p.-group C := pnat_dvd pC (pnat_id p_pr).
+have [_ dv_p] := primeP p_pr; case/pred2P: {dv_p pC}(dv_p _ pC) => [|pC].
+ by move/card1_trivg->; rewrite cent1T setIT trivg_quotient pgroup1.
+have le_oGC := logn_quotient_cent_cyclic_pgroup pgC.
+rewrite /pgroup -partn_eq1 ?cardG_gt0 // -dvdn1 p_part pfactor_dvdn // logn1.
+by rewrite (leq_trans (le_oGC _)) ?prime_cyclic // pC ?(pfactorK 1).
+Qed.
+
+End InnerAutCyclicPgroup.
+
+Section PcoreDef.
+
+(* A functor needs to quantify over the finGroupType just beore the set. *)
+
+Variables (pi : nat_pred) (gT : finGroupType) (A : {set gT}).
+
+Definition pcore := \bigcap_(G | [max G | pi.-subgroup(A) G]) G.
+
+Canonical pcore_group : {group gT} := Eval hnf in [group of pcore].
+
+End PcoreDef.
+
+Arguments Scope pcore [_ nat_scope group_scope].
+Arguments Scope pcore_group [_ nat_scope Group_scope].
+Notation "''O_' pi ( G )" := (pcore pi G)
+ (at level 8, pi at level 2, format "''O_' pi ( G )") : group_scope.
+Notation "''O_' pi ( G )" := (pcore_group pi G) : Group_scope.
+
+Section PseriesDefs.
+
+Variables (pis : seq nat_pred) (gT : finGroupType) (A : {set gT}).
+
+Definition pcore_mod pi B := coset B @*^-1 'O_pi(A / B).
+Canonical pcore_mod_group pi B : {group gT} :=
+ Eval hnf in [group of pcore_mod pi B].
+
+Definition pseries := foldr pcore_mod 1 (rev pis).
+
+Lemma pseries_group_set : group_set pseries.
+Proof. rewrite /pseries; case: rev => [|pi1 pi1']; exact: groupP. Qed.
+
+Canonical pseries_group : {group gT} := group pseries_group_set.
+
+End PseriesDefs.
+
+Arguments Scope pseries [_ seq_scope group_scope].
+Local Notation ConsPred p := (@Cons nat_pred p%N) (only parsing).
+Notation "''O_{' p1 , .. , pn } ( A )" :=
+ (pseries (ConsPred p1 .. (ConsPred pn [::]) ..) A)
+ (at level 8, format "''O_{' p1 , .. , pn } ( A )") : group_scope.
+Notation "''O_{' p1 , .. , pn } ( A )" :=
+ (pseries_group (ConsPred p1 .. (ConsPred pn [::]) ..) A) : Group_scope.
+
+Section PCoreProps.
+
+Variables (pi : nat_pred) (gT : finGroupType).
+Implicit Types (A : {set gT}) (G H M K : {group gT}).
+
+Lemma pcore_psubgroup G : pi.-subgroup(G) 'O_pi(G).
+Proof.
+have [M maxM _]: {M | [max M | pi.-subgroup(G) M] & 1%G \subset M}.
+ by apply: maxgroup_exists; rewrite /psubgroup sub1G pgroup1.
+have sOM: 'O_pi(G) \subset M by exact: bigcap_inf.
+have /andP[piM sMG] := maxgroupp maxM.
+by rewrite /psubgroup (pgroupS sOM) // (subset_trans sOM).
+Qed.
+
+Lemma pcore_pgroup G : pi.-group 'O_pi(G).
+Proof. by case/andP: (pcore_psubgroup G). Qed.
+
+Lemma pcore_sub G : 'O_pi(G) \subset G.
+Proof. by case/andP: (pcore_psubgroup G). Qed.
+
+Lemma pcore_sub_Hall G H : pi.-Hall(G) H -> 'O_pi(G) \subset H.
+Proof. by move/Hall_max=> maxH; exact: bigcap_inf. Qed.
+
+Lemma pcore_max G H : pi.-group H -> H <| G -> H \subset 'O_pi(G).
+Proof.
+move=> piH nHG; apply/bigcapsP=> M maxM.
+exact: normal_sub_max_pgroup piH nHG.
+Qed.
+
+Lemma pcore_pgroup_id G : pi.-group G -> 'O_pi(G) = G.
+Proof. by move=> piG; apply/eqP; rewrite eqEsubset pcore_sub pcore_max. Qed.
+
+Lemma pcore_normal G : 'O_pi(G) <| G.
+Proof.
+rewrite /(_ <| G) pcore_sub; apply/subsetP=> x Gx.
+rewrite inE; apply/bigcapsP=> M maxM; rewrite sub_conjg.
+by apply: bigcap_inf; apply: max_pgroupJ; rewrite ?groupV.
+Qed.
+
+Lemma normal_Hall_pcore H G : pi.-Hall(G) H -> H <| G -> 'O_pi(G) = H.
+Proof.
+move=> hallH nHG; apply/eqP.
+rewrite eqEsubset (sub_normal_Hall hallH) ?pcore_sub ?pcore_pgroup //=.
+by rewrite pcore_max //= (pHall_pgroup hallH).
+Qed.
+
+Lemma eq_Hall_pcore G H :
+ pi.-Hall(G) 'O_pi(G) -> pi.-Hall(G) H -> H :=: 'O_pi(G).
+Proof.
+move=> hallGpi hallH.
+exact: uniq_normal_Hall (pcore_normal G) (Hall_max hallH).
+Qed.
+
+Lemma sub_Hall_pcore G K :
+ pi.-Hall(G) 'O_pi(G) -> K \subset G -> (K \subset 'O_pi(G)) = pi.-group K.
+Proof. by move=> hallGpi; exact: sub_normal_Hall (pcore_normal G). Qed.
+
+Lemma mem_Hall_pcore G x :
+ pi.-Hall(G) 'O_pi(G) -> x \in G -> (x \in 'O_pi(G)) = pi.-elt x.
+Proof. move=> hallGpi; exact: mem_normal_Hall (pcore_normal G). Qed.
+
+Lemma sdprod_Hall_pcoreP H G :
+ pi.-Hall(G) 'O_pi(G) -> reflect ('O_pi(G) ><| H = G) (pi^'.-Hall(G) H).
+Proof.
+move=> hallGpi; rewrite -(compl_pHall H hallGpi) complgC.
+exact: sdprod_normal_complP (pcore_normal G).
+Qed.
+
+Lemma sdprod_pcore_HallP H G :
+ pi^'.-Hall(G) H -> reflect ('O_pi(G) ><| H = G) (pi.-Hall(G) 'O_pi(G)).
+Proof. exact: sdprod_normal_p'HallP (pcore_normal G). Qed.
+
+Lemma pcoreJ G x : 'O_pi(G :^ x) = 'O_pi(G) :^ x.
+Proof.
+apply/eqP; rewrite eqEsubset -sub_conjgV.
+rewrite !pcore_max ?pgroupJ ?pcore_pgroup ?normalJ ?pcore_normal //.
+by rewrite -(normalJ _ _ x) conjsgKV pcore_normal.
+Qed.
+
+End PCoreProps.
+
+Section MorphPcore.
+
+Implicit Types (pi : nat_pred) (gT rT : finGroupType).
+
+Lemma morphim_pcore pi : GFunctor.pcontinuous (pcore pi).
+Proof.
+move=> gT rT D G f; apply/bigcapsP=> M /normal_sub_max_pgroup; apply.
+ by rewrite morphim_pgroup ?pcore_pgroup.
+apply: morphim_normal; exact: pcore_normal.
+Qed.
+
+Lemma pcoreS pi gT (G H : {group gT}) :
+ H \subset G -> H :&: 'O_pi(G) \subset 'O_pi(H).
+Proof.
+move=> sHG; rewrite -{2}(setIidPl sHG).
+do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; exact: morphim_pcore.
+Qed.
+
+Canonical pcore_igFun pi := [igFun by pcore_sub pi & morphim_pcore pi].
+Canonical pcore_gFun pi := [gFun by morphim_pcore pi].
+Canonical pcore_pgFun pi := [pgFun by morphim_pcore pi].
+
+Lemma pcore_char pi gT (G : {group gT}) : 'O_pi(G) \char G.
+Proof. exact: gFchar. Qed.
+
+Section PcoreMod.
+
+Variable F : GFunctor.pmap.
+
+Lemma pcore_mod_sub pi gT (G : {group gT}) : pcore_mod G pi (F _ G) \subset G.
+Proof.
+have nFD := gFnorm F G; rewrite sub_morphpre_im ?pcore_sub //=.
+ by rewrite ker_coset_prim subIset // gen_subG gFsub.
+by apply: subset_trans (pcore_sub _ _) _; apply: morphimS.
+Qed.
+
+Lemma quotient_pcore_mod pi gT (G : {group gT}) (B : {set gT}) :
+ pcore_mod G pi B / B = 'O_pi(G / B).
+Proof.
+apply: morphpreK; apply: subset_trans (pcore_sub _ _) _.
+by rewrite /= /quotient -morphimIdom morphimS ?subsetIl.
+Qed.
+
+Lemma morphim_pcore_mod pi gT rT (D G : {group gT}) (f : {morphism D >-> rT}) :
+ f @* pcore_mod G pi (F _ G) \subset pcore_mod (f @* G) pi (F _ (f @* G)).
+Proof.
+have sDF: D :&: G \subset 'dom (coset (F _ G)).
+ by rewrite setIC subIset ?gFnorm.
+have sDFf: D :&: G \subset 'dom (coset (F _ (f @* G)) \o f).
+ by rewrite -sub_morphim_pre ?subsetIl // morphimIdom gFnorm.
+pose K := 'ker (restrm sDFf (coset (F _ (f @* G)) \o f)).
+have sFK: 'ker (restrm sDF (coset (F _ G))) \subset K.
+ rewrite /K !ker_restrm ker_comp /= subsetI subsetIl /= -setIA.
+ rewrite -sub_morphim_pre ?subsetIl //.
+ by rewrite morphimIdom !ker_coset (setIidPr _) ?pmorphimF ?gFsub.
+have sOF := pcore_sub pi (G / F _ G); have sDD: D :&: G \subset D :&: G by [].
+rewrite -sub_morphim_pre -?quotientE; last first.
+ by apply: subset_trans (gFnorm F _); rewrite morphimS ?pcore_mod_sub.
+suffices im_fact (H : {group gT}) : F _ G \subset H -> H \subset G ->
+ factm sFK sDD @* (H / F _ G) = f @* H / F _ (f @* G).
+- rewrite -2?im_fact ?pcore_mod_sub ?gFsub //;
+ try by rewrite -{1}[F _ G]ker_coset morphpreS ?sub1G.
+ by rewrite quotient_pcore_mod morphim_pcore.
+move=> sFH sHG; rewrite -(morphimIdom _ (H / _)) /= {2}morphim_restrm setIid.
+rewrite -morphimIG ?ker_coset //.
+rewrite -(morphim_restrm sDF) morphim_factm morphim_restrm.
+by rewrite morphim_comp -quotientE -setIA morphimIdom (setIidPr _).
+Qed.
+
+Lemma pcore_mod_res pi gT rT (D : {group gT}) (f : {morphism D >-> rT}) :
+ f @* pcore_mod D pi (F _ D) \subset pcore_mod (f @* D) pi (F _ (f @* D)).
+Proof. exact: morphim_pcore_mod. Qed.
+
+Lemma pcore_mod1 pi gT (G : {group gT}) : pcore_mod G pi 1 = 'O_pi(G).
+Proof.
+rewrite /pcore_mod; have inj1 := coset1_injm gT; rewrite -injmF ?norms1 //.
+by rewrite -(morphim_invmE inj1) morphim_invm ?norms1.
+Qed.
+
+End PcoreMod.
+
+Lemma pseries_rcons pi pis gT (A : {set gT}) :
+ pseries (rcons pis pi) A = pcore_mod A pi (pseries pis A).
+Proof. by rewrite /pseries rev_rcons. Qed.
+
+Lemma pseries_subfun pis :
+ GFunctor.closed (pseries pis) /\ GFunctor.pcontinuous (pseries pis).
+Proof.
+elim/last_ind: pis => [|pis pi [sFpi fFpi]].
+ by split=> [gT G | gT rT D G f]; rewrite (sub1G, morphim1).
+pose fF := [gFun by fFpi : GFunctor.continuous [igFun by sFpi & fFpi]].
+pose F := [pgFun by fFpi : GFunctor.hereditary fF].
+split=> [gT G | gT rT D G f]; rewrite !pseries_rcons ?(pcore_mod_sub F) //.
+exact: (morphim_pcore_mod F).
+Qed.
+
+Lemma pseries_sub pis : GFunctor.closed (pseries pis).
+Proof. by case: (pseries_subfun pis). Qed.
+
+Lemma morphim_pseries pis : GFunctor.pcontinuous (pseries pis).
+Proof. by case: (pseries_subfun pis). Qed.
+
+Lemma pseriesS pis : GFunctor.hereditary (pseries pis).
+Proof. exact: (morphim_pseries pis). Qed.
+
+Canonical pseries_igFun pis := [igFun by pseries_sub pis & morphim_pseries pis].
+Canonical pseries_gFun pis := [gFun by morphim_pseries pis].
+Canonical pseries_pgFun pis := [pgFun by morphim_pseries pis].
+
+Lemma pseries_char pis gT (G : {group gT}) : pseries pis G \char G.
+Proof. exact: gFchar. Qed.
+
+Lemma pseries_normal pis gT (G : {group gT}) : pseries pis G <| G.
+Proof. exact: gFnormal. Qed.
+
+Lemma pseriesJ pis gT (G : {group gT}) x :
+ pseries pis (G :^ x) = pseries pis G :^ x.
+Proof.
+rewrite -{1}(setIid G) -morphim_conj -(injmF _ (injm_conj G x)) //=.
+by rewrite morphim_conj (setIidPr (pseries_sub _ _)).
+Qed.
+
+Lemma pseries1 pi gT (G : {group gT}) : 'O_{pi}(G) = 'O_pi(G).
+Proof. exact: pcore_mod1. Qed.
+
+Lemma pseries_pop pi pis gT (G : {group gT}) :
+ 'O_pi(G) = 1 -> pseries (pi :: pis) G = pseries pis G.
+Proof.
+by move=> OG1; rewrite /pseries rev_cons -cats1 foldr_cat /= pcore_mod1 OG1.
+Qed.
+
+Lemma pseries_pop2 pi1 pi2 gT (G : {group gT}) :
+ 'O_pi1(G) = 1 -> 'O_{pi1, pi2}(G) = 'O_pi2(G).
+Proof. by move/pseries_pop->; exact: pseries1. Qed.
+
+Lemma pseries_sub_catl pi1s pi2s gT (G : {group gT}) :
+ pseries pi1s G \subset pseries (pi1s ++ pi2s) G.
+Proof.
+elim/last_ind: pi2s => [|pi pis IHpi]; rewrite ?cats0 // -rcons_cat.
+by rewrite pseries_rcons; apply: subset_trans IHpi _; rewrite sub_cosetpre.
+Qed.
+
+Lemma quotient_pseries pis pi gT (G : {group gT}) :
+ pseries (rcons pis pi) G / pseries pis G = 'O_pi(G / pseries pis G).
+Proof. by rewrite pseries_rcons quotient_pcore_mod. Qed.
+
+Lemma pseries_norm2 pi1s pi2s gT (G : {group gT}) :
+ pseries pi2s G \subset 'N(pseries pi1s G).
+Proof.
+apply: subset_trans (normal_norm (pseries_normal pi1s G)); exact: pseries_sub.
+Qed.
+
+Lemma pseries_sub_catr pi1s pi2s gT (G : {group gT}) :
+ pseries pi2s G \subset pseries (pi1s ++ pi2s) G.
+Proof.
+elim: pi1s => //= pi1 pi1s /subset_trans; apply.
+elim/last_ind: {pi1s pi2s}(_ ++ _) => [|pis pi IHpi]; first exact: sub1G.
+rewrite -rcons_cons (pseries_rcons _ (pi1 :: pis)).
+rewrite -sub_morphim_pre ?pseries_norm2 //.
+apply: pcore_max; last by rewrite morphim_normal ?pseries_normal.
+have: pi.-group (pseries (rcons pis pi) G / pseries pis G).
+ by rewrite quotient_pseries pcore_pgroup.
+by apply: pnat_dvd; rewrite !card_quotient ?pseries_norm2 // indexgS.
+Qed.
+
+Lemma quotient_pseries2 pi1 pi2 gT (G : {group gT}) :
+ 'O_{pi1, pi2}(G) / 'O_pi1(G) = 'O_pi2(G / 'O_pi1(G)).
+Proof. by rewrite -pseries1 -quotient_pseries. Qed.
+
+Lemma quotient_pseries_cat pi1s pi2s gT (G : {group gT}) :
+ pseries (pi1s ++ pi2s) G / pseries pi1s G
+ = pseries pi2s (G / pseries pi1s G).
+Proof.
+elim/last_ind: pi2s => [|pi2s pi IHpi]; first by rewrite cats0 trivg_quotient.
+have psN := pseries_normal _ G; set K := pseries _ G.
+case: (third_isom (pseries_sub_catl pi1s pi2s G) (psN _)) => //= f inj_f im_f.
+have nH2H: pseries pi2s (G / K) <| pseries (pi1s ++ rcons pi2s pi) G / K.
+ rewrite -IHpi morphim_normal // -cats1 catA.
+ by apply/andP; rewrite pseries_sub_catl pseries_norm2.
+apply: (quotient_inj nH2H).
+ by apply/andP; rewrite /= -cats1 pseries_sub_catl pseries_norm2.
+rewrite /= quotient_pseries /= -IHpi -rcons_cat.
+rewrite -[G / _ / _](morphim_invm inj_f) //= {2}im_f //.
+rewrite -(@injmF [igFun of pcore pi]) /= ?injm_invm ?im_f // -quotient_pseries.
+by rewrite -im_f ?morphim_invm ?morphimS ?normal_sub.
+Qed.
+
+Lemma pseries_catl_id pi1s pi2s gT (G : {group gT}) :
+ pseries pi1s (pseries (pi1s ++ pi2s) G) = pseries pi1s G.
+Proof.
+elim/last_ind: pi1s => [//|pi1s pi IHpi] in pi2s *.
+apply: (@quotient_inj _ (pseries_group pi1s G)).
+- rewrite /= -(IHpi (pi :: pi2s)) cat_rcons /(_ <| _) pseries_norm2.
+ by rewrite -cats1 pseries_sub_catl.
+- by rewrite /= /(_ <| _) pseries_norm2 -cats1 pseries_sub_catl.
+rewrite /= cat_rcons -(IHpi (pi :: pi2s)) {1}quotient_pseries IHpi.
+apply/eqP; rewrite quotient_pseries eqEsubset !pcore_max ?pcore_pgroup //=.
+ rewrite -quotient_pseries morphim_normal // /(_ <| _) pseries_norm2.
+ by rewrite -cat_rcons pseries_sub_catl.
+by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?gFnormal.
+Qed.
+
+Lemma pseries_char_catl pi1s pi2s gT (G : {group gT}) :
+ pseries pi1s G \char pseries (pi1s ++ pi2s) G.
+Proof. by rewrite -(pseries_catl_id pi1s pi2s G) pseries_char. Qed.
+
+Lemma pseries_catr_id pi1s pi2s gT (G : {group gT}) :
+ pseries pi2s (pseries (pi1s ++ pi2s) G) = pseries pi2s G.
+Proof.
+elim/last_ind: pi2s => [//|pi2s pi IHpi] in G *.
+have Epis: pseries pi2s (pseries (pi1s ++ rcons pi2s pi) G) = pseries pi2s G.
+ by rewrite -cats1 catA -2!IHpi pseries_catl_id.
+apply: (@quotient_inj _ (pseries_group pi2s G)).
+- by rewrite /= -Epis /(_ <| _) pseries_norm2 -cats1 pseries_sub_catl.
+- by rewrite /= /(_ <| _) pseries_norm2 -cats1 pseries_sub_catl.
+rewrite /= -Epis {1}quotient_pseries Epis quotient_pseries.
+apply/eqP; rewrite eqEsubset !pcore_max ?pcore_pgroup //=.
+ rewrite -quotient_pseries morphim_normal // /(_ <| _) pseries_norm2.
+ by rewrite pseries_sub_catr.
+apply: char_normal_trans (pcore_char pi _) (morphim_normal _ _).
+exact: pseries_normal.
+Qed.
+
+Lemma pseries_char_catr pi1s pi2s gT (G : {group gT}) :
+ pseries pi2s G \char pseries (pi1s ++ pi2s) G.
+Proof. by rewrite -(pseries_catr_id pi1s pi2s G) pseries_char. Qed.
+
+Lemma pcore_modp pi gT (G H : {group gT}) :
+ H <| G -> pi.-group H -> pcore_mod G pi H = 'O_pi(G).
+Proof.
+move=> nsHG piH; apply/eqP; rewrite eqEsubset andbC.
+have nHG := normal_norm nsHG; have sOG := subset_trans (pcore_sub pi _).
+rewrite -sub_morphim_pre ?(sOG, morphim_pcore) // pcore_max //.
+ rewrite -(pquotient_pgroup piH) ?subsetIl //.
+ by rewrite quotient_pcore_mod pcore_pgroup.
+by rewrite -{2}(quotientGK nsHG) morphpre_normal ?pcore_normal ?sOG ?morphimS.
+Qed.
+
+Lemma pquotient_pcore pi gT (G H : {group gT}) :
+ H <| G -> pi.-group H -> 'O_pi(G / H) = 'O_pi(G) / H.
+Proof. by move=> nsHG piH; rewrite -quotient_pcore_mod pcore_modp. Qed.
+
+Lemma trivg_pcore_quotient pi gT (G : {group gT}) : 'O_pi(G / 'O_pi(G)) = 1.
+Proof.
+by rewrite pquotient_pcore ?pcore_normal ?pcore_pgroup // trivg_quotient.
+Qed.
+
+Lemma pseries_rcons_id pis pi gT (G : {group gT}) :
+ pseries (rcons (rcons pis pi) pi) G = pseries (rcons pis pi) G.
+Proof.
+apply/eqP; rewrite -!cats1 eqEsubset pseries_sub_catl andbT -catA.
+rewrite -(quotientSGK _ (pseries_sub_catl _ _ _)) ?pseries_norm2 //.
+rewrite !quotient_pseries_cat -quotient_sub1 ?pseries_norm2 //.
+by rewrite quotient_pseries_cat /= !pseries1 trivg_pcore_quotient.
+Qed.
+
+End MorphPcore.
+
+Section EqPcore.
+
+Variables gT : finGroupType.
+Implicit Types (pi rho : nat_pred) (G H : {group gT}).
+
+Lemma sub_in_pcore pi rho G :
+ {in \pi(G), {subset pi <= rho}} -> 'O_pi(G) \subset 'O_rho(G).
+Proof.
+move=> pi_sub_rho; rewrite pcore_max ?pcore_normal //.
+apply: sub_in_pnat (pcore_pgroup _ _) => p.
+move/(piSg (pcore_sub _ _)); exact: pi_sub_rho.
+Qed.
+
+Lemma sub_pcore pi rho G : {subset pi <= rho} -> 'O_pi(G) \subset 'O_rho(G).
+Proof. by move=> pi_sub_rho; exact: sub_in_pcore (in1W pi_sub_rho). Qed.
+
+Lemma eq_in_pcore pi rho G : {in \pi(G), pi =i rho} -> 'O_pi(G) = 'O_rho(G).
+Proof.
+move=> eq_pi_rho; apply/eqP; rewrite eqEsubset.
+by rewrite !sub_in_pcore // => p /eq_pi_rho->.
+Qed.
+
+Lemma eq_pcore pi rho G : pi =i rho -> 'O_pi(G) = 'O_rho(G).
+Proof. by move=> eq_pi_rho; exact: eq_in_pcore (in1W eq_pi_rho). Qed.
+
+Lemma pcoreNK pi G : 'O_pi^'^'(G) = 'O_pi(G).
+Proof. by apply: eq_pcore; exact: negnK. Qed.
+
+Lemma eq_p'core pi rho G : pi =i rho -> 'O_pi^'(G) = 'O_rho^'(G).
+Proof. by move/eq_negn; exact: eq_pcore. Qed.
+
+Lemma sdprod_Hall_p'coreP pi H G :
+ pi^'.-Hall(G) 'O_pi^'(G) -> reflect ('O_pi^'(G) ><| H = G) (pi.-Hall(G) H).
+Proof. by rewrite -(pHallNK pi G H); exact: sdprod_Hall_pcoreP. Qed.
+
+Lemma sdprod_p'core_HallP pi H G :
+ pi.-Hall(G) H -> reflect ('O_pi^'(G) ><| H = G) (pi^'.-Hall(G) 'O_pi^'(G)).
+Proof. by rewrite -(pHallNK pi G H); exact: sdprod_pcore_HallP. Qed.
+
+Lemma pcoreI pi rho G : 'O_[predI pi & rho](G) = 'O_pi('O_rho(G)).
+Proof.
+apply/eqP; rewrite eqEsubset !pcore_max //.
+- rewrite /pgroup pnatI [pnat _ _]pcore_pgroup.
+ exact: pgroupS (pcore_sub _ _) (pcore_pgroup _ _).
+- exact: char_normal_trans (pcore_char _ _) (pcore_normal _ _).
+- by apply: sub_in_pnat (pcore_pgroup _ _) => p _ /andP[].
+apply/andP; split; first by apply: sub_pcore => p /andP[].
+by rewrite (subset_trans (pcore_sub _ _)) ?gFnorm.
+Qed.
+
+Lemma bigcap_p'core pi G :
+ G :&: \bigcap_(p < #|G|.+1 | (p : nat) \in pi) 'O_p^'(G) = 'O_pi^'(G).
+Proof.
+apply/eqP; rewrite eqEsubset subsetI pcore_sub pcore_max /=.
+- by apply/bigcapsP=> p pi_p; apply: sub_pcore => r; apply: contraNneq => ->.
+- apply/pgroupP=> q q_pr qGpi'; apply: contraL (eqxx q) => /= pi_q.
+ apply: (pgroupP (pcore_pgroup q^' G)) => //.
+ have qG: q %| #|G| by rewrite (dvdn_trans qGpi') // cardSg ?subsetIl.
+ have ltqG: q < #|G|.+1 by rewrite ltnS dvdn_leq.
+ rewrite (dvdn_trans qGpi') ?cardSg ?subIset //= orbC.
+ by rewrite (bigcap_inf (Ordinal ltqG)).
+rewrite /normal subsetIl normsI ?normG // norms_bigcap //.
+by apply/bigcapsP => p _; exact: gFnorm.
+Qed.
+
+Lemma coprime_pcoreC (rT : finGroupType) pi G (R : {group rT}) :
+ coprime #|'O_pi(G)| #|'O_pi^'(R)|.
+Proof. exact: pnat_coprime (pcore_pgroup _ _) (pcore_pgroup _ _). Qed.
+
+Lemma TI_pcoreC pi G H : 'O_pi(G) :&: 'O_pi^'(H) = 1.
+Proof. by rewrite coprime_TIg ?coprime_pcoreC. Qed.
+
+Lemma pcore_setI_normal pi G H : H <| G -> 'O_pi(G) :&: H = 'O_pi(H).
+Proof.
+move=> nsHG; apply/eqP; rewrite eqEsubset subsetI pcore_sub.
+rewrite !pcore_max ?(pgroupS (subsetIl _ H)) ?pcore_pgroup //=.
+ exact: char_normal_trans (pcore_char pi H) nsHG.
+by rewrite setIC (normalGI (normal_sub nsHG)) ?pcore_normal.
+Qed.
+
+End EqPcore.
+
+Implicit Arguments sdprod_Hall_pcoreP [gT pi G H].
+Implicit Arguments sdprod_Hall_p'coreP [gT pi G H].
+
+Section Injm.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}).
+Hypothesis injf : 'injm f.
+Implicit Types (A : {set aT}) (G H : {group aT}).
+
+Lemma injm_pgroup pi A : A \subset D -> pi.-group (f @* A) = pi.-group A.
+Proof. by move=> sAD; rewrite /pgroup card_injm. Qed.
+
+Lemma injm_pelt pi x : x \in D -> pi.-elt (f x) = pi.-elt x.
+Proof. by move=> Dx; rewrite /p_elt order_injm. Qed.
+
+Lemma injm_pHall pi G H :
+ G \subset D -> H \subset D -> pi.-Hall(f @* G) (f @* H) = pi.-Hall(G) H.
+Proof. by move=> sGD sGH; rewrite !pHallE injmSK ?card_injm. Qed.
+
+Lemma injm_pcore pi G : G \subset D -> f @* 'O_pi(G) = 'O_pi(f @* G).
+Proof. exact: injmF. Qed.
+
+Lemma injm_pseries pis G :
+ G \subset D -> f @* pseries pis G = pseries pis (f @* G).
+Proof. exact: injmF. Qed.
+
+End Injm.
+
+Section Isog.
+
+Variables (aT rT : finGroupType) (G : {group aT}) (H : {group rT}).
+
+Lemma isog_pgroup pi : G \isog H -> pi.-group G = pi.-group H.
+Proof. by move=> isoGH; rewrite /pgroup (card_isog isoGH). Qed.
+
+Lemma isog_pcore pi : G \isog H -> 'O_pi(G) \isog 'O_pi(H).
+Proof. exact: gFisog. Qed.
+
+Lemma isog_pseries pis : G \isog H -> pseries pis G \isog pseries pis H.
+Proof. exact: gFisog. Qed.
+
+End Isog.
diff --git a/mathcomp/solvable/primitive_action.v b/mathcomp/solvable/primitive_action.v
new file mode 100644
index 0000000..7e2075d
--- /dev/null
+++ b/mathcomp/solvable/primitive_action.v
@@ -0,0 +1,347 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat.
+Require Import div seq fintype tuple finset.
+Require Import fingroup action gseries.
+
+(******************************************************************************)
+(* n-transitive and primitive actions: *)
+(* [primitive A, on S | to] <=> *)
+(* A acts on S in a primitive manner, i.e., A is transitive on S and *)
+(* A does not act on any nontrivial partition of S. *)
+(* imprimitivity_system A to S Q <=> *)
+(* Q is a non-trivial primitivity system for the action of A on S via *)
+(* to, i.e., Q is a non-trivial partiiton of S on which A acts. *)
+(* to * n == in the %act scope, the total action induced by the total *)
+(* action to on n.-tuples. via n_act to n. *)
+(* n.-dtuple S == the set of n-tuples with distinct values in S. *)
+(* [transitive^n A, on S | to] <=> *)
+(* A is n-transitive on S, i.e., A is transitive on n.-dtuple S *)
+(* == the set of n-tuples with distinct values in S. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+Section PrimitiveDef.
+
+Variables (aT : finGroupType) (sT : finType).
+Variables (A : {set aT}) (S : {set sT}) (to : {action aT &-> sT}).
+
+Definition imprimitivity_system Q :=
+ [&& partition Q S, [acts A, on Q | to^*] & 1 < #|Q| < #|S|].
+
+Definition primitive :=
+ [transitive A, on S | to] && ~~ [exists Q, imprimitivity_system Q].
+
+End PrimitiveDef.
+
+Arguments Scope imprimitivity_system
+ [_ _ group_scope group_scope action_scope group_scope].
+Arguments Scope primitive [_ _ group_scope group_scope action_scope].
+
+Notation "[ 'primitive' A , 'on' S | to ]" := (primitive A S to)
+ (at level 0, format "[ 'primitive' A , 'on' S | to ]") : form_scope.
+
+Prenex Implicits imprimitivity_system.
+
+Section Primitive.
+
+Variables (aT : finGroupType) (sT : finType).
+Variables (G : {group aT}) (to : {action aT &-> sT}) (S : {set sT}).
+
+Lemma trans_prim_astab x :
+ x \in S -> [transitive G, on S | to] ->
+ [primitive G, on S | to] = maximal_eq 'C_G[x | to] G.
+Proof.
+move=> Sx trG; rewrite /primitive trG negb_exists.
+apply/forallP/maximal_eqP=> /= [primG | [_ maxCx] Q].
+ split=> [|H sCH sHG]; first exact: subsetIl.
+ pose X := orbit to H x; pose Q := orbit (to^*)%act G X.
+ have Xx: x \in X by exact: orbit_refl.
+ have defH: 'N_(G)(X | to) = H.
+ have trH: [transitive H, on X | to] by apply/imsetP; exists x.
+ have sHN: H \subset 'N_G(X | to) by rewrite subsetI sHG atrans_acts.
+ move/(subgroup_transitiveP Xx sHN): (trH) => /= <-.
+ by rewrite mulSGid //= setIAC subIset ?sCH.
+ apply/imsetP; exists x => //; apply/eqP.
+ by rewrite eqEsubset imsetS // acts_sub_orbit ?subsetIr.
+ have [|/proper_card oCH] := eqVproper sCH; [by left | right].
+ apply/eqP; rewrite eqEcard sHG leqNgt.
+ apply: contra {primG}(primG Q) => oHG; apply/and3P; split; last first.
+ - rewrite card_orbit astab1_set defH -(@ltn_pmul2l #|H|) ?Lagrange // muln1.
+ rewrite oHG -(@ltn_pmul2l #|H|) ?Lagrange // -(card_orbit_stab to G x).
+ by rewrite -(atransP trG x Sx) mulnC card_orbit ltn_pmul2r.
+ - by apply/actsP=> a Ga Y; apply: orbit_transr; exact: mem_orbit.
+ apply/and3P; split; last 1 first.
+ - rewrite orbit_sym; apply/imsetP=> [[a _]] /= defX.
+ by rewrite defX /setact imset0 inE in Xx.
+ - apply/eqP/setP=> y; apply/bigcupP/idP=> [[_ /imsetP[a Ga ->]] | Sy].
+ case/imsetP=> _ /imsetP[b Hb ->] ->.
+ by rewrite !(actsP (atrans_acts trG)) //; exact: subsetP Hb.
+ case: (atransP2 trG Sx Sy) => a Ga ->.
+ by exists ((to^*)%act X a); apply: mem_imset; rewrite // orbit_refl.
+ apply/trivIsetP=> _ _ /imsetP[a Ga ->] /imsetP[b Gb ->].
+ apply: contraR => /exists_inP[_ /imsetP[_ /imsetP[a1 Ha1 ->] ->]].
+ case/imsetP=> _ /imsetP[b1 Hb1 ->] /(canLR (actK _ _)) /(canLR (actK _ _)).
+ rewrite -(canF_eq (actKV _ _)) -!actM (sameP eqP astab1P) => /astab1P Cab.
+ rewrite astab1_set (subsetP (subsetIr G _)) //= defH.
+ rewrite -(groupMr _ (groupVr Hb1)) -mulgA -(groupMl _ Ha1).
+ by rewrite (subsetP sCH) // inE Cab !groupM ?groupV // (subsetP sHG).
+apply/and3P=> [[/and3P[/eqP defS tIQ ntQ]]]; set sto := (to^*)%act => actQ.
+rewrite !ltnNge -negb_or => /orP[].
+pose X := pblock Q x; have Xx: x \in X by rewrite mem_pblock defS.
+have QX: X \in Q by rewrite pblock_mem ?defS.
+have toX Y a: Y \in Q -> a \in G -> to x a \in Y -> sto X a = Y.
+ move=> QY Ga Yxa; rewrite -(contraNeq (trivIsetP tIQ Y (sto X a) _ _)) //.
+ by rewrite (actsP actQ).
+ by apply/existsP; exists (to x a); rewrite /= Yxa; apply: mem_imset.
+have defQ: Q = orbit (to^*)%act G X.
+ apply/eqP; rewrite eqEsubset andbC acts_sub_orbit // QX.
+ apply/subsetP=> Y QY.
+ have /set0Pn[y Yy]: Y != set0 by apply: contraNneq ntQ => <-.
+ have Sy: y \in S by rewrite -defS; apply/bigcupP; exists Y.
+ have [a Ga def_y] := atransP2 trG Sx Sy.
+ by apply/imsetP; exists a; rewrite // (toX Y) // -def_y.
+rewrite defQ card_orbit; case: (maxCx 'C_G[X | sto]%G) => /= [||->|->].
+- apply/subsetP=> a /setIP[Ga cxa]; rewrite inE Ga /=.
+ by apply/astab1P; rewrite (toX X) // (astab1P cxa).
+- exact: subsetIl.
+- by right; rewrite -card_orbit (atransP trG).
+by left; rewrite indexgg.
+Qed.
+
+Lemma prim_trans_norm (H : {group aT}) :
+ [primitive G, on S | to] -> H <| G ->
+ H \subset 'C_G(S | to) \/ [transitive H, on S | to].
+Proof.
+move=> primG /andP[sHG nHG]; rewrite subsetI sHG.
+have [trG _] := andP primG; have [x Sx defS] := imsetP trG.
+move: primG; rewrite (trans_prim_astab Sx) // => /maximal_eqP[_].
+case/(_ ('C_G[x | to] <*> H)%G) => /= [||cxH|]; first exact: joing_subl.
+- by rewrite join_subG subsetIl.
+- have{cxH} cxH: H \subset 'C_G[x | to] by rewrite -cxH joing_subr.
+ rewrite subsetI sHG /= in cxH; left; apply/subsetP=> a Ha.
+ apply/astabP=> y Sy; have [b Gb ->] := atransP2 trG Sx Sy.
+ rewrite actCJV [to x (a ^ _)](astab1P _) ?(subsetP cxH) //.
+ by rewrite -mem_conjg (normsP nHG).
+rewrite norm_joinEl 1?subIset ?nHG //.
+by move/(subgroup_transitiveP Sx sHG trG); right.
+Qed.
+
+End Primitive.
+
+Section NactionDef.
+
+Variables (gT : finGroupType) (sT : finType).
+Variables (to : {action gT &-> sT}) (n : nat).
+
+Definition n_act (t : n.-tuple sT) a := [tuple of map (to^~ a) t].
+
+Fact n_act_is_action : is_action setT n_act.
+Proof.
+by apply: is_total_action => [t|t a b]; apply: eq_from_tnth => i;
+ rewrite !tnth_map ?act1 ?actM.
+Qed.
+
+Canonical n_act_action := Action n_act_is_action.
+
+End NactionDef.
+
+Notation "to * n" := (n_act_action to n) : action_scope.
+
+Section NTransitive.
+
+Variables (gT : finGroupType) (sT : finType).
+Variables (n : nat) (A : {set gT}) (S : {set sT}) (to : {action gT &-> sT}).
+
+Definition dtuple_on := [set t : n.-tuple sT | uniq t & t \subset S].
+Definition ntransitive := [transitive A, on dtuple_on | to * n].
+
+Lemma dtuple_onP t :
+ reflect (injective (tnth t) /\ forall i, tnth t i \in S) (t \in dtuple_on).
+Proof.
+rewrite inE subset_all -map_tnth_enum.
+case: (uniq _) / (injectiveP (tnth t)) => f_inj; last by right; case.
+rewrite -[all _ _]negbK -has_predC has_map has_predC negbK /=.
+by apply: (iffP allP) => [Sf|[]//]; split=> // i; rewrite Sf ?mem_enum.
+Qed.
+
+Lemma n_act_dtuple t a :
+ a \in 'N(S | to) -> t \in dtuple_on -> n_act to t a \in dtuple_on.
+Proof.
+move/astabsP=> toSa /dtuple_onP[t_inj St]; apply/dtuple_onP.
+split=> [i j | i]; rewrite !tnth_map ?[_ \in S]toSa //.
+by move/act_inj; exact: t_inj.
+Qed.
+
+End NTransitive.
+
+Arguments Scope dtuple_on [_ nat_scope group_scope].
+Arguments Scope ntransitive
+ [_ _ nat_scope group_scope group_scope action_scope].
+Implicit Arguments n_act [gT sT n].
+
+Notation "n .-dtuple ( S )" := (dtuple_on n S)
+ (at level 8, format "n .-dtuple ( S )") : set_scope.
+
+Notation "[ 'transitive' ^ n A , 'on' S | to ]" := (ntransitive n A S to)
+ (at level 0, n at level 8,
+ format "[ 'transitive' ^ n A , 'on' S | to ]") : form_scope.
+
+Section NTransitveProp.
+
+Variables (gT : finGroupType) (sT : finType).
+Variables (to : {action gT &-> sT}) (G : {group gT}) (S : {set sT}).
+
+Lemma card_uniq_tuple n (t : n.-tuple sT) : uniq t -> #|t| = n.
+Proof. by move/card_uniqP->; exact: size_tuple. Qed.
+
+Lemma n_act0 (t : 0.-tuple sT) a : n_act to t a = [tuple].
+Proof. exact: tuple0. Qed.
+
+Lemma dtuple_on_add n x (t : n.-tuple sT) :
+ ([tuple of x :: t] \in n.+1.-dtuple(S)) =
+ [&& x \in S, x \notin t & t \in n.-dtuple(S)].
+Proof. by rewrite !inE memtE !subset_all -!andbA; do !bool_congr. Qed.
+
+Lemma dtuple_on_add_D1 n x (t : n.-tuple sT) :
+ ([tuple of x :: t] \in n.+1.-dtuple(S))
+ = (x \in S) && (t \in n.-dtuple(S :\ x)).
+Proof.
+rewrite dtuple_on_add !inE (andbCA (~~ _)); do 2!congr (_ && _).
+rewrite -!(eq_subset (in_set (mem t))) setDE setIC subsetI; congr (_ && _).
+by rewrite -setCS setCK sub1set !inE.
+Qed.
+
+Lemma dtuple_on_subset n (S1 S2 : {set sT}) t :
+ S1 \subset S2 -> t \in n.-dtuple(S1) -> t \in n.-dtuple(S2).
+Proof. by move=> sS12; rewrite !inE => /andP[-> /subset_trans]; exact. Qed.
+
+Lemma n_act_add n x (t : n.-tuple sT) a :
+ n_act to [tuple of x :: t] a = [tuple of to x a :: n_act to t a].
+Proof. exact: val_inj. Qed.
+
+Lemma ntransitive0 : [transitive^0 G, on S | to].
+Proof.
+have dt0: [tuple] \in 0.-dtuple(S) by rewrite inE memtE subset_all.
+apply/imsetP; exists [tuple of Nil sT] => //.
+by apply/setP=> x; rewrite [x]tuple0 orbit_refl.
+Qed.
+
+Lemma ntransitive_weak k m :
+ k <= m -> [transitive^m G, on S | to] -> [transitive^k G, on S | to].
+Proof.
+move/subnKC <-; rewrite addnC; elim: {m}(m - k) => // m IHm.
+rewrite addSn => tr_m1; apply: IHm; move: {m k}(m + k) tr_m1 => m tr_m1.
+have ext_t t: t \in dtuple_on m S ->
+ exists x, [tuple of x :: t] \in m.+1.-dtuple(S).
+- move=> dt.
+ have [sSt | /subsetPn[x Sx ntx]] := boolP (S \subset t); last first.
+ by exists x; rewrite dtuple_on_add andbA /= Sx ntx.
+ case/imsetP: tr_m1 dt => t1; rewrite !inE => /andP[Ut1 St1] _ /andP[Ut _].
+ have /subset_leq_card := subset_trans St1 sSt.
+ by rewrite !card_uniq_tuple // ltnn.
+case/imsetP: (tr_m1); case/tupleP=> [x t]; rewrite dtuple_on_add.
+case/and3P=> Sx ntx dt; set xt := [tuple of _] => tr_xt.
+apply/imsetP; exists t => //.
+apply/setP=> u; apply/idP/imsetP=> [du | [a Ga ->{u}]].
+ case: (ext_t u du) => y; rewrite tr_xt.
+ by case/imsetP=> a Ga [_ def_u]; exists a => //; exact: val_inj.
+have: n_act to xt a \in dtuple_on _ S by rewrite tr_xt mem_imset.
+by rewrite n_act_add dtuple_on_add; case/and3P.
+Qed.
+
+Lemma ntransitive1 m :
+ 0 < m -> [transitive^m G, on S | to] -> [transitive G, on S | to].
+Proof.
+have trdom1 x: ([tuple x] \in 1.-dtuple(S)) = (x \in S).
+ by rewrite dtuple_on_add !inE memtE subset_all andbT.
+move=> m_gt0 /(ntransitive_weak m_gt0) {m m_gt0}.
+case/imsetP; case/tupleP=> x t0; rewrite {t0}(tuple0 t0) trdom1 => Sx trx.
+apply/imsetP; exists x => //; apply/setP=> y; rewrite -trdom1 trx.
+apply/imsetP/imsetP=> [[a ? [->]]|[a ? ->]]; exists a => //; exact: val_inj.
+Qed.
+
+Lemma ntransitive_primitive m :
+ 1 < m -> [transitive^m G, on S | to] -> [primitive G, on S | to].
+Proof.
+move=> lt1m /(ntransitive_weak lt1m) {m lt1m}tr2G.
+have trG: [transitive G, on S | to] by exact: ntransitive1 tr2G.
+have [x Sx _]:= imsetP trG; rewrite (trans_prim_astab Sx trG).
+apply/maximal_eqP; split=> [|H]; first exact: subsetIl; rewrite subEproper.
+case/predU1P; first by [left]; case/andP=> sCH /subsetPn[a Ha nCa] sHG.
+right; rewrite -(subgroup_transitiveP Sx sHG trG _) ?mulSGid //.
+have actH := subset_trans sHG (atrans_acts trG).
+pose y := to x a; have Sy: y \in S by rewrite (actsP actH).
+have{nCa} yx: y != x by rewrite inE (sameP astab1P eqP) (subsetP sHG) in nCa.
+apply/imsetP; exists y => //; apply/eqP.
+rewrite eqEsubset acts_sub_orbit // Sy andbT; apply/subsetP=> z Sz.
+have [-> | zx] := eqVneq z x; first by rewrite orbit_sym mem_orbit.
+pose ty := [tuple y; x]; pose tz := [tuple z; x].
+have [Sty Stz]: ty \in 2.-dtuple(S) /\ tz \in 2.-dtuple(S).
+ rewrite !inE !memtE !subset_all /= !mem_seq1 !andbT; split; exact/and3P.
+case: (atransP2 tr2G Sty Stz) => b Gb [->] /esym/astab1P cxb.
+by rewrite mem_orbit // (subsetP sCH) // inE Gb.
+Qed.
+
+End NTransitveProp.
+
+Section NTransitveProp1.
+
+Variables (gT : finGroupType) (sT : finType).
+Variables (to : {action gT &-> sT}) (G : {group gT}) (S : {set sT}).
+
+(* This is the forward implication of Aschbacher (15.12).1 *)
+Theorem stab_ntransitive m x :
+ 0 < m -> x \in S -> [transitive^m.+1 G, on S | to] ->
+ [transitive^m 'C_G[x | to], on S :\ x | to].
+Proof.
+move=> m_gt0 Sx Gtr; have sSxS: S :\ x \subset S by rewrite subsetDl.
+case: (imsetP Gtr); case/tupleP=> x1 t1; rewrite dtuple_on_add.
+case/and3P=> Sx1 nt1x1 dt1 trt1; have Gtr1 := ntransitive1 (ltn0Sn _) Gtr.
+case: (atransP2 Gtr1 Sx1 Sx) => // a Ga x1ax.
+pose t := n_act to t1 a.
+have dxt: [tuple of x :: t] \in m.+1.-dtuple(S).
+ rewrite trt1 x1ax; apply/imsetP; exists a => //; exact: val_inj.
+apply/imsetP; exists t; first by rewrite dtuple_on_add_D1 Sx in dxt.
+apply/setP=> t2; apply/idP/imsetP => [dt2|[b]].
+ have: [tuple of x :: t2] \in dtuple_on _ S by rewrite dtuple_on_add_D1 Sx.
+ case/(atransP2 Gtr dxt)=> b Gb [xbx tbt2].
+ exists b; [rewrite inE Gb; exact/astab1P | exact: val_inj].
+case/setIP=> Gb /astab1P xbx ->{t2}.
+rewrite n_act_dtuple //; last by rewrite dtuple_on_add_D1 Sx in dxt.
+apply/astabsP=> y; rewrite !inE -{1}xbx (inj_eq (act_inj _ _)).
+by rewrite (actsP (atrans_acts Gtr1)).
+Qed.
+
+(* This is the converse implication of Aschbacher (15.12).1 *)
+Theorem stab_ntransitiveI m x :
+ x \in S -> [transitive G, on S | to] ->
+ [transitive^m 'C_G[x | to], on S :\ x | to] ->
+ [transitive^m.+1 G, on S | to].
+Proof.
+move=> Sx Gtr Gntr.
+have t_to_x t: t \in m.+1.-dtuple(S) ->
+ exists2 a, a \in G & exists2 t', t' \in m.-dtuple(S :\ x)
+ & t = n_act to [tuple of x :: t'] a.
+- case/tupleP: t => y t St.
+ have Sy: y \in S by rewrite dtuple_on_add_D1 in St; case/andP: St.
+ rewrite -(atransP Gtr _ Sy) in Sx; case/imsetP: Sx => a Ga toya.
+ exists a^-1; first exact: groupVr.
+ exists (n_act to t a); last by rewrite n_act_add toya !actK.
+ move/(n_act_dtuple (subsetP (atrans_acts Gtr) a Ga)): St.
+ by rewrite n_act_add -toya dtuple_on_add_D1 => /andP[].
+case: (imsetP Gntr) => t dt S_tG; pose xt := [tuple of x :: t].
+have dxt: xt \in m.+1.-dtuple(S) by rewrite dtuple_on_add_D1 Sx.
+apply/imsetP; exists xt => //; apply/setP=> t2.
+apply/esym; apply/imsetP/idP=> [[a Ga ->] | ].
+ by apply: n_act_dtuple; rewrite // (subsetP (atrans_acts Gtr)).
+case/t_to_x=> a2 Ga2 [t2']; rewrite S_tG.
+case/imsetP=> a /setIP[Ga /astab1P toxa] -> -> {t2 t2'}.
+by exists (a * a2); rewrite (groupM, actM) //= !n_act_add toxa.
+Qed.
+
+End NTransitveProp1.
diff --git a/mathcomp/solvable/sylow.v b/mathcomp/solvable/sylow.v
new file mode 100644
index 0000000..02ab37a
--- /dev/null
+++ b/mathcomp/solvable/sylow.v
@@ -0,0 +1,661 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype prime.
+Require Import bigop finset fingroup morphism automorphism quotient action.
+Require Import cyclic gproduct commutator pgroup center nilpotent.
+
+(******************************************************************************)
+(* The Sylow theorem and its consequences, including the Frattini argument, *)
+(* the nilpotence of p-groups, and the Baer-Suzuki theorem. *)
+(* This file also defines: *)
+(* Zgroup G == G is a Z-group, i.e., has only cyclic Sylow p-subgroups. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Import GroupScope.
+
+(* The mod p lemma for the action of p-groups. *)
+Section ModP.
+
+Variable (aT : finGroupType) (sT : finType) (D : {group aT}).
+Variable to : action D sT.
+
+Lemma pgroup_fix_mod (p : nat) (G : {group aT}) (S : {set sT}) :
+ p.-group G -> [acts G, on S | to] -> #|S| = #|'Fix_(S | to)(G)| %[mod p].
+Proof.
+move=> pG nSG; have sGD: G \subset D := acts_dom nSG.
+apply/eqP; rewrite -(cardsID 'Fix_to(G)) eqn_mod_dvd (leq_addr, addKn) //.
+have: [acts G, on S :\: 'Fix_to(G) | to]; last move/acts_sum_card_orbit <-.
+ rewrite actsD // -(setIidPr sGD); apply: subset_trans (acts_subnorm_fix _ _).
+ by rewrite setIS ?normG.
+apply: dvdn_sum => _ /imsetP[x /setDP[_ nfx] ->].
+have [k oGx]: {k | #|orbit to G x| = (p ^ k)%N}.
+ by apply: p_natP; apply: pnat_dvd pG; rewrite card_orbit_in ?dvdn_indexg.
+case: k oGx => [/card_orbit1 fix_x | k ->]; last by rewrite expnS dvdn_mulr.
+by case/afixP: nfx => a Ga; apply/set1P; rewrite -fix_x mem_orbit.
+Qed.
+
+End ModP.
+
+Section ModularGroupAction.
+
+Variables (aT rT : finGroupType) (D : {group aT}) (R : {group rT}).
+Variables (to : groupAction D R) (p : nat).
+Implicit Types (G H : {group aT}) (M : {group rT}).
+
+Lemma nontrivial_gacent_pgroup G M :
+ p.-group G -> p.-group M -> {acts G, on group M | to} ->
+ M :!=: 1 -> 'C_(M | to)(G) :!=: 1.
+Proof.
+move=> pG pM [nMG sMR] ntM; have [p_pr p_dv_M _] := pgroup_pdiv pM ntM.
+rewrite -cardG_gt1 (leq_trans (prime_gt1 p_pr)) 1?dvdn_leq ?cardG_gt0 //= /dvdn.
+by rewrite gacentE ?(acts_dom nMG) // setIA (setIidPl sMR) -pgroup_fix_mod.
+Qed.
+
+Lemma pcore_sub_astab_irr G M :
+ p.-group M -> M \subset R -> acts_irreducibly G M to ->
+ 'O_p(G) \subset 'C_G(M | to).
+Proof.
+move=> pM sMR /mingroupP[/andP[ntM nMG] minM].
+have [sGpG nGpG]:= andP (pcore_normal p G).
+have sGD := acts_dom nMG; have sGpD := subset_trans sGpG sGD.
+rewrite subsetI sGpG -gacentC //=; apply/setIidPl; apply: minM (subsetIl _ _).
+rewrite nontrivial_gacent_pgroup ?pcore_pgroup //=; last first.
+ by split; rewrite ?(subset_trans sGpG).
+by apply: subset_trans (acts_subnorm_subgacent sGpD nMG); rewrite subsetI subxx.
+Qed.
+
+Lemma pcore_faithful_irr_act G M :
+ p.-group M -> M \subset R -> acts_irreducibly G M to ->
+ [faithful G, on M | to] ->
+ 'O_p(G) = 1.
+Proof.
+move=> pM sMR irrG ffulG; apply/trivgP; apply: subset_trans ffulG.
+exact: pcore_sub_astab_irr.
+Qed.
+
+End ModularGroupAction.
+
+Section Sylow.
+
+Variables (p : nat) (gT : finGroupType) (G : {group gT}).
+Implicit Types P Q H K : {group gT}.
+
+Theorem Sylow's_theorem :
+ [/\ forall P, [max P | p.-subgroup(G) P] = p.-Sylow(G) P,
+ [transitive G, on 'Syl_p(G) | 'JG],
+ forall P, p.-Sylow(G) P -> #|'Syl_p(G)| = #|G : 'N_G(P)|
+ & prime p -> #|'Syl_p(G)| %% p = 1%N].
+Proof.
+pose maxp A P := [max P | p.-subgroup(A) P]; pose S := [set P | maxp G P].
+pose oG := orbit 'JG%act G.
+have actS: [acts G, on S | 'JG].
+ apply/subsetP=> x Gx; rewrite 3!inE; apply/subsetP=> P; rewrite 3!inE.
+ exact: max_pgroupJ.
+have S_pG P: P \in S -> P \subset G /\ p.-group P.
+ by rewrite inE => /maxgroupp/andP[].
+have SmaxN P Q: Q \in S -> Q \subset 'N(P) -> maxp 'N_G(P) Q.
+ rewrite inE => /maxgroupP[/andP[sQG pQ] maxQ] nPQ.
+ apply/maxgroupP; rewrite /psubgroup subsetI sQG nPQ.
+ by split=> // R; rewrite subsetI -andbA andbCA => /andP[_]; exact: maxQ.
+have nrmG P: P \subset G -> P <| 'N_G(P).
+ by move=> sPG; rewrite /normal subsetIr subsetI sPG normG.
+have sylS P: P \in S -> p.-Sylow('N_G(P)) P.
+ move=> S_P; have [sPG pP] := S_pG P S_P.
+ by rewrite normal_max_pgroup_Hall ?nrmG //; apply: SmaxN; rewrite ?normG.
+have{SmaxN} defCS P: P \in S -> 'Fix_(S |'JG)(P) = [set P].
+ move=> S_P; apply/setP=> Q; rewrite {1}in_setI {1}afixJG.
+ apply/andP/set1P=> [[S_Q nQP]|->{Q}]; last by rewrite normG.
+ apply/esym/val_inj; case: (S_pG Q) => //= sQG _.
+ by apply: uniq_normal_Hall (SmaxN Q _ _ _) => //=; rewrite ?sylS ?nrmG.
+have{defCS} oG_mod: {in S &, forall P Q, #|oG P| = (Q \in oG P) %[mod p]}.
+ move=> P Q S_P S_Q; have [sQG pQ] := S_pG _ S_Q.
+ have soP_S: oG P \subset S by rewrite acts_sub_orbit.
+ have /pgroup_fix_mod-> //: [acts Q, on oG P | 'JG].
+ apply/actsP=> x /(subsetP sQG) Gx R; apply: orbit_transr.
+ exact: mem_orbit.
+ rewrite -{1}(setIidPl soP_S) -setIA defCS // (cardsD1 Q) setDE.
+ by rewrite -setIA setICr setI0 cards0 addn0 inE set11 andbT.
+have [P S_P]: exists P, P \in S.
+ have: p.-subgroup(G) 1 by rewrite /psubgroup sub1G pgroup1.
+ by case/(@maxgroup_exists _ (p.-subgroup(G))) => P; exists P; rewrite inE.
+have trS: [transitive G, on S | 'JG].
+ apply/imsetP; exists P => //; apply/eqP.
+ rewrite eqEsubset andbC acts_sub_orbit // S_P; apply/subsetP=> Q S_Q.
+ have:= S_P; rewrite inE => /maxgroupP[/andP[_ pP]].
+ have [-> max1 | ntP _] := eqVneq P 1%G.
+ move/andP/max1: (S_pG _ S_Q) => Q1.
+ by rewrite (group_inj (Q1 (sub1G Q))) orbit_refl.
+ have:= oG_mod _ _ S_P S_P; rewrite (oG_mod _ Q) // orbit_refl.
+ have p_gt1: p > 1 by apply: prime_gt1; case/pgroup_pdiv: pP.
+ by case: (Q \in oG P) => //; rewrite mod0n modn_small.
+have oS1: prime p -> #|S| %% p = 1%N.
+ move/prime_gt1 => p_gt1.
+ by rewrite -(atransP trS P S_P) (oG_mod P P) // orbit_refl modn_small.
+have oSiN Q: Q \in S -> #|S| = #|G : 'N_G(Q)|.
+ by move=> S_Q; rewrite -(atransP trS Q S_Q) card_orbit astab1JG.
+have sylP: p.-Sylow(G) P.
+ rewrite pHallE; case: (S_pG P) => // -> /= pP.
+ case p_pr: (prime p); last first.
+ rewrite p_part lognE p_pr /= -trivg_card1; apply/idPn=> ntP.
+ by case/pgroup_pdiv: pP p_pr => // ->.
+ rewrite -(LagrangeI G 'N(P)) /= mulnC partnM ?cardG_gt0 // part_p'nat.
+ by rewrite mul1n (card_Hall (sylS P S_P)).
+ by rewrite p'natE // -indexgI -oSiN // /dvdn oS1.
+have eqS Q: maxp G Q = p.-Sylow(G) Q.
+ apply/idP/idP=> [S_Q|]; last exact: Hall_max.
+ have{S_Q} S_Q: Q \in S by rewrite inE.
+ rewrite pHallE -(card_Hall sylP); case: (S_pG Q) => // -> _ /=.
+ by case: (atransP2 trS S_P S_Q) => x _ ->; rewrite cardJg.
+have ->: 'Syl_p(G) = S by apply/setP=> Q; rewrite 2!inE.
+by split=> // Q sylQ; rewrite -oSiN ?inE ?eqS.
+Qed.
+
+Lemma max_pgroup_Sylow P : [max P | p.-subgroup(G) P] = p.-Sylow(G) P.
+Proof. by case Sylow's_theorem. Qed.
+
+Lemma Sylow_superset Q :
+ Q \subset G -> p.-group Q -> {P : {group gT} | p.-Sylow(G) P & Q \subset P}.
+Proof.
+move=> sQG pQ.
+have [|P] := @maxgroup_exists _ (p.-subgroup(G)) Q; first exact/andP.
+by rewrite max_pgroup_Sylow; exists P.
+Qed.
+
+Lemma Sylow_exists : {P : {group gT} | p.-Sylow(G) P}.
+Proof. by case: (Sylow_superset (sub1G G) (pgroup1 _ p)) => P; exists P. Qed.
+
+Lemma Syl_trans : [transitive G, on 'Syl_p(G) | 'JG].
+Proof. by case Sylow's_theorem. Qed.
+
+Lemma Sylow_trans P Q :
+ p.-Sylow(G) P -> p.-Sylow(G) Q -> exists2 x, x \in G & Q :=: P :^ x.
+Proof.
+move=> sylP sylQ; have:= (atransP2 Syl_trans) P Q; rewrite !inE.
+by case=> // x Gx ->; exists x.
+Qed.
+
+Lemma Sylow_subJ P Q :
+ p.-Sylow(G) P -> Q \subset G -> p.-group Q ->
+ exists2 x, x \in G & Q \subset P :^ x.
+Proof.
+move=> sylP sQG pQ; have [Px sylPx] := Sylow_superset sQG pQ.
+by have [x Gx ->] := Sylow_trans sylP sylPx; exists x.
+Qed.
+
+Lemma Sylow_Jsub P Q :
+ p.-Sylow(G) P -> Q \subset G -> p.-group Q ->
+ exists2 x, x \in G & Q :^ x \subset P.
+Proof.
+move=> sylP sQG pQ; have [x Gx] := Sylow_subJ sylP sQG pQ.
+by exists x^-1; rewrite (groupV, sub_conjgV).
+Qed.
+
+Lemma card_Syl P : p.-Sylow(G) P -> #|'Syl_p(G)| = #|G : 'N_G(P)|.
+Proof. by case: Sylow's_theorem P. Qed.
+
+Lemma card_Syl_dvd : #|'Syl_p(G)| %| #|G|.
+Proof. by case Sylow_exists => P /card_Syl->; exact: dvdn_indexg. Qed.
+
+Lemma card_Syl_mod : prime p -> #|'Syl_p(G)| %% p = 1%N.
+Proof. by case Sylow's_theorem. Qed.
+
+Lemma Frattini_arg H P : G <| H -> p.-Sylow(G) P -> G * 'N_H(P) = H.
+Proof.
+case/andP=> sGH nGH sylP; rewrite -normC ?subIset ?nGH ?orbT // -astab1JG.
+move/subgroup_transitiveP: Syl_trans => ->; rewrite ?inE //.
+apply/imsetP; exists P; rewrite ?inE //.
+apply/eqP; rewrite eqEsubset -{1}((atransP Syl_trans) P) ?inE // imsetS //=.
+by apply/subsetP=> _ /imsetP[x Hx ->]; rewrite inE -(normsP nGH x Hx) pHallJ2.
+Qed.
+
+End Sylow.
+
+Section MoreSylow.
+
+Variables (gT : finGroupType) (p : nat).
+Implicit Types G H P : {group gT}.
+
+Lemma Sylow_setI_normal G H P :
+ G <| H -> p.-Sylow(H) P -> p.-Sylow(G) (G :&: P).
+Proof.
+case/normalP=> sGH nGH sylP; have [Q sylQ] := Sylow_exists p G.
+have /maxgroupP[/andP[sQG pQ] maxQ] := Hall_max sylQ.
+have [R sylR sQR] := Sylow_superset (subset_trans sQG sGH) pQ.
+have [[x Hx ->] pR] := (Sylow_trans sylR sylP, pHall_pgroup sylR).
+rewrite -(nGH x Hx) -conjIg pHallJ2.
+have /maxQ-> //: Q \subset G :&: R by rewrite subsetI sQG.
+by rewrite /psubgroup subsetIl (pgroupS _ pR) ?subsetIr.
+Qed.
+
+Lemma normal_sylowP G :
+ reflect (exists2 P : {group gT}, p.-Sylow(G) P & P <| G)
+ (#|'Syl_p(G)| == 1%N).
+Proof.
+apply: (iffP idP) => [syl1 | [P sylP nPG]]; last first.
+ by rewrite (card_Syl sylP) (setIidPl _) (indexgg, normal_norm).
+have [P sylP] := Sylow_exists p G; exists P => //.
+rewrite /normal (pHall_sub sylP); apply/setIidPl; apply/eqP.
+rewrite eqEcard subsetIl -(LagrangeI G 'N(P)) -indexgI /=.
+by rewrite -(card_Syl sylP) (eqP syl1) muln1.
+Qed.
+
+Lemma trivg_center_pgroup P : p.-group P -> 'Z(P) = 1 -> P :=: 1.
+Proof.
+move=> pP Z1; apply/eqP/idPn=> ntP.
+have{ntP} [p_pr p_dv_P _] := pgroup_pdiv pP ntP.
+suff: p %| #|'Z(P)| by rewrite Z1 cards1 gtnNdvd ?prime_gt1.
+by rewrite /center /dvdn -afixJ -pgroup_fix_mod // astabsJ normG.
+Qed.
+
+Lemma p2group_abelian P : p.-group P -> logn p #|P| <= 2 -> abelian P.
+Proof.
+move=> pP lePp2; pose Z := 'Z(P); have sZP: Z \subset P := center_sub P.
+case: (eqVneq Z 1); first by move/(trivg_center_pgroup pP)->; exact: abelian1.
+case/(pgroup_pdiv (pgroupS sZP pP)) => p_pr _ [k oZ].
+apply: cyclic_center_factor_abelian.
+case: (eqVneq (P / Z) 1) => [-> |]; first exact: cyclic1.
+have pPq := quotient_pgroup 'Z(P) pP; case/(pgroup_pdiv pPq) => _ _ [j oPq].
+rewrite prime_cyclic // oPq; case: j oPq lePp2 => //= j.
+rewrite card_quotient ?gfunctor.gFnorm //.
+by rewrite -(Lagrange sZP) lognM // => ->; rewrite oZ !pfactorK ?addnS.
+Qed.
+
+Lemma card_p2group_abelian P : prime p -> #|P| = (p ^ 2)%N -> abelian P.
+Proof.
+move=> primep oP; have pP: p.-group P by rewrite /pgroup oP pnat_exp pnat_id.
+by rewrite (p2group_abelian pP) // oP pfactorK.
+Qed.
+
+Lemma Sylow_transversal_gen (T : {set {group gT}}) G :
+ (forall P, P \in T -> P \subset G) ->
+ (forall p, p \in \pi(G) -> exists2 P, P \in T & p.-Sylow(G) P) ->
+ << \bigcup_(P in T) P >> = G.
+Proof.
+move=> G_T T_G; apply/eqP; rewrite eqEcard gen_subG.
+apply/andP; split; first exact/bigcupsP.
+apply: dvdn_leq (cardG_gt0 _) _; apply/dvdn_partP=> // q /T_G[P T_P sylP].
+by rewrite -(card_Hall sylP); apply: cardSg; rewrite sub_gen // bigcup_sup.
+Qed.
+
+Lemma Sylow_gen G : <<\bigcup_(P : {group gT} | Sylow G P) P>> = G.
+Proof.
+set T := [set P : {group gT} | Sylow G P].
+rewrite -{2}(@Sylow_transversal_gen T G) => [|P | q _].
+- by congr <<_>>; apply: eq_bigl => P; rewrite inE.
+- by rewrite inE => /and3P[].
+by case: (Sylow_exists q G) => P sylP; exists P; rewrite // inE (p_Sylow sylP).
+Qed.
+
+End MoreSylow.
+
+Section SomeHall.
+
+Variable gT : finGroupType.
+Implicit Types (p : nat) (pi : nat_pred) (G H K P R : {group gT}).
+
+Lemma Hall_pJsub p pi G H P :
+ pi.-Hall(G) H -> p \in pi -> P \subset G -> p.-group P ->
+ exists2 x, x \in G & P :^ x \subset H.
+Proof.
+move=> hallH pi_p sPG pP.
+have [S sylS] := Sylow_exists p H; have sylS_G := subHall_Sylow hallH pi_p sylS.
+have [x Gx sPxS] := Sylow_Jsub sylS_G sPG pP; exists x => //.
+exact: subset_trans sPxS (pHall_sub sylS).
+Qed.
+
+Lemma Hall_psubJ p pi G H P :
+ pi.-Hall(G) H -> p \in pi -> P \subset G -> p.-group P ->
+ exists2 x, x \in G & P \subset H :^ x.
+Proof.
+move=> hallH pi_p sPG pP; have [x Gx sPxH] := Hall_pJsub hallH pi_p sPG pP.
+by exists x^-1; rewrite ?groupV -?sub_conjg.
+Qed.
+
+Lemma Hall_setI_normal pi G K H :
+ K <| G -> pi.-Hall(G) H -> pi.-Hall(K) (H :&: K).
+Proof.
+move=> nsKG hallH; have [sHG piH _] := and3P hallH.
+have [sHK_H sHK_K] := (subsetIl H K, subsetIr H K).
+rewrite pHallE sHK_K /= -(part_pnat_id (pgroupS sHK_H piH)); apply/eqP.
+rewrite (widen_partn _ (subset_leq_card sHK_K)); apply: eq_bigr => p pi_p.
+have [P sylP] := Sylow_exists p H.
+have sylPK := Sylow_setI_normal nsKG (subHall_Sylow hallH pi_p sylP).
+rewrite -!p_part -(card_Hall sylPK); symmetry; apply: card_Hall.
+by rewrite (pHall_subl _ sHK_K) //= setIC setSI ?(pHall_sub sylP).
+Qed.
+
+Lemma coprime_mulG_setI_norm H G K R :
+ K * R = G -> G \subset 'N(H) -> coprime #|K| #|R| ->
+ (K :&: H) * (R :&: H) = G :&: H.
+Proof.
+move=> defG nHG coKR; apply/eqP; rewrite eqEcard mulG_subG /= -defG.
+rewrite !setSI ?mulG_subl ?mulG_subr //=.
+rewrite coprime_cardMg ?(coKR, coprimeSg (subsetIl _ _), coprime_sym) //=.
+pose pi := \pi(K); have piK: pi.-group K by exact: pgroup_pi.
+have pi'R: pi^'.-group R by rewrite /pgroup -coprime_pi' /=.
+have [hallK hallR] := coprime_mulpG_Hall defG piK pi'R.
+have nsHG: H :&: G <| G by rewrite /normal subsetIr normsI ?normG.
+rewrite -!(setIC H) defG -(partnC pi (cardG_gt0 _)).
+rewrite -(card_Hall (Hall_setI_normal nsHG hallR)) /= setICA.
+rewrite -(card_Hall (Hall_setI_normal nsHG hallK)) /= setICA.
+by rewrite -defG (setIidPl (mulG_subl _ _)) (setIidPl (mulG_subr _ _)).
+Qed.
+
+End SomeHall.
+
+Section Nilpotent.
+
+Variable gT : finGroupType.
+Implicit Types (G H K P L : {group gT}) (p q : nat).
+
+Lemma pgroup_nil p P : p.-group P -> nilpotent P.
+Proof.
+move: {2}_.+1 (ltnSn #|P|) => n.
+elim: n gT P => // n IHn pT P; rewrite ltnS=> lePn pP.
+have [Z1 | ntZ] := eqVneq 'Z(P) 1.
+ by rewrite (trivg_center_pgroup pP Z1) nilpotent1.
+rewrite -quotient_center_nil IHn ?morphim_pgroup // (leq_trans _ lePn) //.
+rewrite card_quotient ?normal_norm ?center_normal // -divgS ?subsetIl //.
+by rewrite ltn_Pdiv // ltnNge -trivg_card_le1.
+Qed.
+
+Lemma pgroup_sol p P : p.-group P -> solvable P.
+Proof. by move/pgroup_nil; exact: nilpotent_sol. Qed.
+
+Lemma small_nil_class G : nil_class G <= 5 -> nilpotent G.
+Proof.
+move=> leK5; case: (ltnP 5 #|G|) => [lt5G | leG5 {leK5}].
+ by rewrite nilpotent_class (leq_ltn_trans leK5).
+apply: pgroup_nil (pdiv #|G|) _ _; apply/andP; split=> //.
+by case: #|G| leG5 => //; do 5!case=> //.
+Qed.
+
+Lemma nil_class2 G : (nil_class G <= 2) = (G^`(1) \subset 'Z(G)).
+Proof.
+rewrite subsetI der_sub; apply/idP/commG1P=> [clG2 | L3G1].
+ by apply/(lcn_nil_classP 2); rewrite ?small_nil_class ?(leq_trans clG2).
+by apply/(lcn_nil_classP 2) => //; apply/lcnP; exists 2.
+Qed.
+
+Lemma nil_class3 G : (nil_class G <= 3) = ('L_3(G) \subset 'Z(G)).
+Proof.
+rewrite subsetI lcn_sub; apply/idP/commG1P=> [clG3 | L4G1].
+ by apply/(lcn_nil_classP 3); rewrite ?small_nil_class ?(leq_trans clG3).
+by apply/(lcn_nil_classP 3) => //; apply/lcnP; exists 3.
+Qed.
+
+Lemma nilpotent_maxp_normal pi G H :
+ nilpotent G -> [max H | pi.-subgroup(G) H] -> H <| G.
+Proof.
+move=> nilG /maxgroupP[/andP[sHG piH] maxH].
+have nHN: H <| 'N_G(H) by rewrite normal_subnorm.
+have{maxH} hallH: pi.-Hall('N_G(H)) H.
+ apply: normal_max_pgroup_Hall => //; apply/maxgroupP.
+ rewrite /psubgroup normal_sub // piH; split=> // K.
+ by rewrite subsetI -andbA andbCA => /andP[_]; exact: maxH.
+rewrite /normal sHG; apply/setIidPl; symmetry.
+apply: nilpotent_sub_norm; rewrite ?subsetIl ?setIS //=.
+by rewrite char_norms // -{1}(normal_Hall_pcore hallH) // pcore_char.
+Qed.
+
+Lemma nilpotent_Hall_pcore pi G H :
+ nilpotent G -> pi.-Hall(G) H -> H :=: 'O_pi(G).
+Proof.
+move=> nilG hallH; have maxH := Hall_max hallH; apply/eqP.
+rewrite eqEsubset pcore_max ?(pHall_pgroup hallH) //.
+ by rewrite (normal_sub_max_pgroup maxH) ?pcore_pgroup ?pcore_normal.
+exact: nilpotent_maxp_normal maxH.
+Qed.
+
+Lemma nilpotent_pcore_Hall pi G : nilpotent G -> pi.-Hall(G) 'O_pi(G).
+Proof.
+move=> nilG; case: (@maxgroup_exists _ (psubgroup pi G) 1) => [|H maxH _].
+ by rewrite /psubgroup sub1G pgroup1.
+have hallH := normal_max_pgroup_Hall maxH (nilpotent_maxp_normal nilG maxH).
+by rewrite -(nilpotent_Hall_pcore nilG hallH).
+Qed.
+
+Lemma nilpotent_pcoreC pi G : nilpotent G -> 'O_pi(G) \x 'O_pi^'(G) = G.
+Proof.
+move=> nilG; have trO: 'O_pi(G) :&: 'O_pi^'(G) = 1.
+ by apply: coprime_TIg; apply: (@pnat_coprime pi); exact: pcore_pgroup.
+rewrite dprodE //.
+ apply/eqP; rewrite eqEcard mul_subG ?pcore_sub // (TI_cardMg trO).
+ by rewrite !(card_Hall (nilpotent_pcore_Hall _ _)) // partnC ?leqnn.
+rewrite (sameP commG1P trivgP) -trO subsetI commg_subl commg_subr.
+by rewrite !(subset_trans (pcore_sub _ _)) ?normal_norm ?pcore_normal.
+Qed.
+
+Lemma sub_nilpotent_cent2 H K G :
+ nilpotent G -> K \subset G -> H \subset G -> coprime #|K| #|H| ->
+ H \subset 'C(K).
+Proof.
+move=> nilG sKG sHG; rewrite coprime_pi' // => p'H.
+have sub_Gp := sub_Hall_pcore (nilpotent_pcore_Hall _ nilG).
+have [_ _ cGpp' _] := dprodP (nilpotent_pcoreC \pi(K) nilG).
+by apply: centSS cGpp'; rewrite sub_Gp ?pgroup_pi.
+Qed.
+
+Lemma pi_center_nilpotent G : nilpotent G -> \pi('Z(G)) = \pi(G).
+Proof.
+move=> nilG; apply/eq_piP => /= p.
+apply/idP/idP=> [|pG]; first exact: (piSg (center_sub _)).
+move: (pG); rewrite !mem_primes !cardG_gt0; case/andP=> p_pr _.
+pose Z := 'O_p(G) :&: 'Z(G); have ntZ: Z != 1.
+ rewrite meet_center_nil ?pcore_normal // trivg_card_le1 -ltnNge.
+ rewrite (card_Hall (nilpotent_pcore_Hall p nilG)) p_part.
+ by rewrite (ltn_exp2l 0 _ (prime_gt1 p_pr)) logn_gt0.
+have pZ: p.-group Z := pgroupS (subsetIl _ _) (pcore_pgroup _ _).
+have{ntZ pZ} [_ pZ _] := pgroup_pdiv pZ ntZ.
+by rewrite p_pr (dvdn_trans pZ) // cardSg ?subsetIr.
+Qed.
+
+Lemma Sylow_subnorm p G P : p.-Sylow('N_G(P)) P = p.-Sylow(G) P.
+Proof.
+apply/idP/idP=> sylP; last first.
+ apply: pHall_subl (subsetIl _ _) (sylP).
+ by rewrite subsetI normG (pHall_sub sylP).
+have [/subsetIP[sPG sPN] pP _] := and3P sylP.
+have [Q sylQ sPQ] := Sylow_superset sPG pP; have [sQG pQ _] := and3P sylQ.
+rewrite -(nilpotent_sub_norm (pgroup_nil pQ) sPQ) {sylQ}//.
+rewrite subEproper eq_sym eqEcard subsetI sPQ sPN dvdn_leq //.
+rewrite -(part_pnat_id (pgroupS (subsetIl _ _) pQ)) (card_Hall sylP).
+by rewrite partn_dvd // cardSg ?setSI.
+Qed.
+
+End Nilpotent.
+
+Lemma nil_class_pgroup (gT : finGroupType) (p : nat) (P : {group gT}) :
+ p.-group P -> nil_class P <= maxn 1 (logn p #|P|).-1.
+Proof.
+move=> pP; move def_c: (nil_class P) => c.
+elim: c => // c IHc in gT P def_c pP *; set e := logn p _.
+have nilP := pgroup_nil pP; have sZP := center_sub P.
+have [e_le2 | e_gt2] := leqP e 2.
+ by rewrite -def_c leq_max nil_class1 (p2group_abelian pP).
+have pPq: p.-group (P / 'Z(P)) by exact: quotient_pgroup.
+rewrite -(subnKC e_gt2) ltnS (leq_trans (IHc _ _ _ pPq)) //.
+ by rewrite nil_class_quotient_center ?def_c.
+rewrite geq_max /= -add1n -leq_subLR -subn1 -subnDA -subSS leq_sub2r //.
+rewrite ltn_log_quotient //= -(setIidPr sZP) meet_center_nil //.
+by rewrite -nil_class0 def_c.
+Qed.
+
+Definition Zgroup (gT : finGroupType) (A : {set gT}) :=
+ [forall (V : {group gT} | Sylow A V), cyclic V].
+
+Section Zgroups.
+
+Variables (gT rT : finGroupType) (D : {group gT}) (f : {morphism D >-> rT}).
+Implicit Types G H K : {group gT}.
+
+Lemma ZgroupS G H : H \subset G -> Zgroup G -> Zgroup H.
+Proof.
+move=> sHG /forallP zgG; apply/forall_inP=> V /SylowP[p p_pr /and3P[sVH]].
+case/(Sylow_superset (subset_trans sVH sHG))=> P sylP sVP _.
+by have:= zgG P; rewrite (p_Sylow sylP); apply: cyclicS.
+Qed.
+
+Lemma morphim_Zgroup G : Zgroup G -> Zgroup (f @* G).
+Proof.
+move=> zgG; wlog sGD: G zgG / G \subset D.
+ by rewrite -morphimIdom; apply; rewrite (ZgroupS _ zgG, subsetIl) ?subsetIr.
+apply/forall_inP=> fV /SylowP[p pr_p sylfV].
+have [P sylP] := Sylow_exists p G.
+have [|z _ ->] := @Sylow_trans p _ _ (f @* P)%G _ _ sylfV.
+ by apply: morphim_pHall (sylP); apply: subset_trans (pHall_sub sylP) sGD.
+by rewrite cyclicJ morphim_cyclic ?(forall_inP zgG) //; apply/SylowP; exists p.
+Qed.
+
+Lemma nil_Zgroup_cyclic G : Zgroup G -> nilpotent G -> cyclic G.
+Proof.
+elim: {G}_.+1 {-2}G (ltnSn #|G|) => // n IHn G; rewrite ltnS => leGn ZgG nilG.
+have [->|[p pr_p pG]] := trivgVpdiv G; first by rewrite -cycle1 cycle_cyclic.
+have /dprodP[_ defG Cpp' _] := nilpotent_pcoreC p nilG.
+have /cyclicP[x def_p]: cyclic 'O_p(G).
+ have:= forallP ZgG 'O_p(G)%G.
+ by rewrite (p_Sylow (nilpotent_pcore_Hall p nilG)).
+have /cyclicP[x' def_p']: cyclic 'O_p^'(G).
+ have sp'G := pcore_sub p^' G.
+ apply: IHn (leq_trans _ leGn) (ZgroupS sp'G _) (nilpotentS sp'G _) => //.
+ rewrite proper_card // properEneq sp'G andbT; case: eqP => //= def_p'.
+ by have:= pcore_pgroup p^' G; rewrite def_p' /pgroup p'natE ?pG.
+apply/cyclicP; exists (x * x'); rewrite -{}defG def_p def_p' cycleM //.
+ by red; rewrite -(centsP Cpp') // (def_p, def_p') cycle_id.
+rewrite /order -def_p -def_p' (@pnat_coprime p) //; exact: pcore_pgroup.
+Qed.
+
+End Zgroups.
+
+Arguments Scope Zgroup [_ group_scope].
+Prenex Implicits Zgroup.
+
+Section NilPGroups.
+
+Variables (p : nat) (gT : finGroupType).
+Implicit Type G P N : {group gT}.
+
+(* B & G 1.22 p.9 *)
+Lemma normal_pgroup r P N :
+ p.-group P -> N <| P -> r <= logn p #|N| ->
+ exists Q : {group gT}, [/\ Q \subset N, Q <| P & #|Q| = (p ^ r)%N].
+Proof.
+elim: r gT P N => [|r IHr] gTr P N pP nNP le_r.
+ by exists (1%G : {group gTr}); rewrite sub1G normal1 cards1.
+have [NZ_1 | ntNZ] := eqVneq (N :&: 'Z(P)) 1.
+ by rewrite (TI_center_nil (pgroup_nil pP)) // cards1 logn1 in le_r.
+have: p.-group (N :&: 'Z(P)) by apply: pgroupS pP; rewrite /= setICA subsetIl.
+case/pgroup_pdiv=> // p_pr /Cauchy[// | z].
+rewrite -cycle_subG !subsetI => /and3P[szN szP cPz] ozp _.
+have{cPz} nzP: P \subset 'N(<[z]>) by rewrite cents_norm // centsC.
+have: N / <[z]> <| P / <[z]> by rewrite morphim_normal.
+case/IHr=> [||Qb [sQNb nQPb]]; first exact: morphim_pgroup.
+ rewrite card_quotient ?(subset_trans (normal_sub nNP)) // -ltnS.
+ apply: (leq_trans le_r); rewrite -(Lagrange szN) [#|_|]ozp.
+ by rewrite lognM // ?prime_gt0 // logn_prime ?eqxx.
+case/(inv_quotientN _): nQPb sQNb => [|Q -> szQ nQP]; first exact/andP.
+have nzQ := subset_trans (normal_sub nQP) nzP.
+rewrite quotientSGK // card_quotient // => sQN izQ.
+by exists Q; split=> //; rewrite expnS -izQ -ozp Lagrange.
+Qed.
+
+Theorem Baer_Suzuki x G :
+ x \in G -> (forall y, y \in G -> p.-group <<[set x; x ^ y]>>) ->
+ x \in 'O_p(G).
+Proof.
+elim: {G}_.+1 {-2}G x (ltnSn #|G|) => // n IHn G x; rewrite ltnS.
+set E := x ^: G => leGn Gx pE.
+have{pE} pE: {in E &, forall x1 x2, p.-group <<[set x1; x2]>>}.
+ move=> _ _ /imsetP[y1 Gy1 ->] /imsetP[y2 Gy2 ->].
+ rewrite -(mulgKV y1 y2) conjgM -2!conjg_set1 -conjUg genJ pgroupJ.
+ by rewrite pE // groupMl ?groupV.
+have sEG: <<E>> \subset G by rewrite gen_subG class_subG.
+have nEG: G \subset 'N(E) by exact: class_norm.
+have Ex: x \in E by exact: class_refl.
+have [P Px sylP]: exists2 P : {group gT}, x \in P & p.-Sylow(<<E>>) P.
+ have sxxE: <<[set x; x]>> \subset <<E>> by rewrite genS // setUid sub1set.
+ have{sxxE} [P sylP sxxP] := Sylow_superset sxxE (pE _ _ Ex Ex).
+ by exists P => //; rewrite (subsetP sxxP) ?mem_gen ?setU11.
+case sEP: (E \subset P).
+ apply: subsetP Ex; rewrite -gen_subG; apply: pcore_max.
+ by apply: pgroupS (pHall_pgroup sylP); rewrite gen_subG.
+ by rewrite /normal gen_subG class_subG // norms_gen.
+pose P_yD D := [pred y in E :\: P | p.-group <<y |: D>>].
+pose P_D := [pred D : {set gT} | D \subset P :&: E & [exists y, P_yD D y]].
+have{Ex Px}: P_D [set x].
+ rewrite /= sub1set inE Px Ex; apply/existsP=> /=.
+ by case/subsetPn: sEP => y Ey Py; exists y; rewrite inE Ey Py pE.
+case/(@maxset_exists _ P_D)=> D /maxsetP[]; rewrite {P_yD P_D}/=.
+rewrite subsetI sub1set -andbA => /and3P[sDP sDE /existsP[y0]].
+set B := _ |: D; rewrite inE -andbA => /and3P[Py0 Ey0 pB] maxD Dx.
+have sDgE: D \subset <<E>> by exact: sub_gen.
+have sDG: D \subset G by exact: subset_trans sEG.
+have sBE: B \subset E by rewrite subUset sub1set Ey0.
+have sBG: <<B>> \subset G by exact: subset_trans (genS _) sEG.
+have sDB: D \subset B by rewrite subsetUr.
+have defD: D :=: P :&: <<B>> :&: E.
+ apply/eqP; rewrite eqEsubset ?subsetI sDP sDE sub_gen //=.
+ apply/setUidPl; apply: maxD; last exact: subsetUl.
+ rewrite subUset subsetI sDP sDE setIAC subsetIl.
+ apply/existsP; exists y0; rewrite inE Py0 Ey0 /= setUA -/B.
+ by rewrite -[<<_>>]joing_idl joingE setKI genGid.
+have nDD: D \subset 'N(D).
+ apply/subsetP=> z Dz; rewrite inE defD.
+ apply/subsetP=> _ /imsetP[y /setIP[PBy Ey] ->].
+ rewrite inE groupJ // ?inE ?(subsetP sDP) ?mem_gen ?setU1r //= memJ_norm //.
+ exact: (subsetP (subset_trans sDG nEG)).
+case nDG: (G \subset 'N(D)).
+ apply: subsetP Dx; rewrite -gen_subG pcore_max ?(pgroupS (genS _) pB) //.
+ by rewrite /normal gen_subG sDG norms_gen.
+have{n leGn IHn nDG} pN: p.-group <<'N_E(D)>>.
+ apply: pgroupS (pcore_pgroup p 'N_G(D)); rewrite gen_subG /=.
+ apply/subsetP=> x1 /setIP[Ex1 Nx1]; apply: IHn => [||y Ny].
+ - apply: leq_trans leGn; rewrite proper_card // /proper subsetIl.
+ by rewrite subsetI nDG andbF.
+ - by rewrite inE Nx1 (subsetP sEG) ?mem_gen.
+ have Ex1y: x1 ^ y \in E.
+ by rewrite -mem_conjgV (normsP nEG) // groupV; case/setIP: Ny.
+ apply: pgroupS (genS _) (pE _ _ Ex1 Ex1y).
+ by apply/subsetP=> u; rewrite !inE.
+have [y1 Ny1 Py1]: exists2 y1, y1 \in 'N_E(D) & y1 \notin P.
+ case sNN: ('N_<<B>>('N_<<B>>(D)) \subset 'N_<<B>>(D)).
+ exists y0 => //; have By0: y0 \in <<B>> by rewrite mem_gen ?setU11.
+ rewrite inE Ey0 -By0 -in_setI.
+ by rewrite -['N__(D)](nilpotent_sub_norm (pgroup_nil pB)) ?subsetIl.
+ case/subsetPn: sNN => z /setIP[Bz NNz]; rewrite inE Bz inE.
+ case/subsetPn=> y; rewrite mem_conjg => Dzy Dy.
+ have:= Dzy; rewrite {1}defD; do 2![case/setIP]=> _ Bzy Ezy.
+ have Ey: y \in E by rewrite -(normsP nEG _ (subsetP sBG z Bz)) mem_conjg.
+ have /setIP[By Ny]: y \in 'N_<<B>>(D).
+ by rewrite -(normP NNz) mem_conjg inE Bzy ?(subsetP nDD).
+ exists y; first by rewrite inE Ey.
+ by rewrite defD 2!inE Ey By !andbT in Dy.
+have [y2 Ny2 Dy2]: exists2 y2, y2 \in 'N_(P :&: E)(D) & y2 \notin D.
+ case sNN: ('N_P('N_P(D)) \subset 'N_P(D)).
+ have [z /= Ez sEzP] := Sylow_Jsub sylP (genS sBE) pB.
+ have Gz: z \in G by exact: subsetP Ez.
+ have /subsetPn[y Bzy Dy]: ~~ (B :^ z \subset D).
+ apply/negP; move/subset_leq_card; rewrite cardJg cardsU1.
+ by rewrite {1}defD 2!inE (negPf Py0) ltnn.
+ exists y => //; apply: subsetP Bzy.
+ rewrite -setIA setICA subsetI sub_conjg (normsP nEG) ?groupV // sBE.
+ have nilP := pgroup_nil (pHall_pgroup sylP).
+ by rewrite -['N__(_)](nilpotent_sub_norm nilP) ?subsetIl // -gen_subG genJ.
+ case/subsetPn: sNN => z /setIP[Pz NNz]; rewrite 2!inE Pz.
+ case/subsetPn=> y Dzy Dy; exists y => //; apply: subsetP Dzy.
+ rewrite -setIA setICA subsetI sub_conjg (normsP nEG) ?groupV //.
+ by rewrite sDE -(normP NNz); rewrite conjSg subsetI sDP.
+ apply: subsetP Pz; exact: (subset_trans (pHall_sub sylP)).
+suff{Dy2} Dy2D: y2 |: D = D by rewrite -Dy2D setU11 in Dy2.
+apply: maxD; last by rewrite subsetUr.
+case/setIP: Ny2 => PEy2 Ny2; case/setIP: Ny1 => Ey1 Ny1.
+rewrite subUset sub1set PEy2 subsetI sDP sDE.
+apply/existsP; exists y1; rewrite inE Ey1 Py1; apply: pgroupS pN.
+rewrite genS // !subUset !sub1set !in_setI Ey1 Ny1.
+by case/setIP: PEy2 => _ ->; rewrite Ny2 subsetI sDE.
+Qed.
+
+End NilPGroups.
diff --git a/mathcomp/solvable/wielandt_fixpoint.v b/mathcomp/solvable/wielandt_fixpoint.v
new file mode 100644
index 0000000..beebc3d
--- /dev/null
+++ b/mathcomp/solvable/wielandt_fixpoint.v
@@ -0,0 +1,651 @@
+(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *)
+Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div.
+Require Import fintype bigop prime binomial finset ssralg fingroup finalg.
+Require Import morphism perm automorphism quotient action commutator gproduct.
+Require Import zmodp cyclic center pgroup gseries nilpotent sylow finalg.
+Require Import finmodule abelian frobenius maximal extremal hall finmodule.
+Require Import matrix mxalgebra mxrepresentation mxabelem BGsection1.
+
+(******************************************************************************)
+(* This file provides the proof of the Wielandt fixpoint order formula, *)
+(* which is a prerequisite for B & G, Section 3 and Peterfalvi, Section 9. *)
+(******************************************************************************)
+
+Set Implicit Arguments.
+Unset Strict Implicit.
+Unset Printing Implicit Defensive.
+
+Local Open Scope ring_scope.
+Import GroupScope GRing.Theory.
+Import FinRing.Theory.
+
+Implicit Types (gT wT : finGroupType) (m n p q : nat).
+
+Lemma coprime_act_abelian_pgroup_structure gT p (A X : {group gT}) :
+ abelian A -> p.-group A -> p^'.-group X -> X \subset 'N(A) ->
+ exists2 s : {set {group gT}},
+ \big[dprod/1]_(B in s) B = A
+ & {in s, forall B : {group gT},
+ [/\ homocyclic B, X \subset 'N(B)
+ & acts_irreducibly X (B / 'Phi(B)) 'Q]}.
+Proof.
+move: {2}_.+1 (ltnSn #|A|) => m.
+elim: m => // m IHm in gT A X *; rewrite ltnS => leAm cAA pA p'X nAX.
+have [n1 eA]: {n | exponent A = p ^ n}%N by apply p_natP; rewrite pnat_exponent.
+have [-> | ntA] := eqsVneq A 1.
+ by exists set0 => [|B]; rewrite ?big_set0 ?inE.
+have [p_pr _ _] := pgroup_pdiv pA ntA; have p_gt1 := prime_gt1 p_pr.
+case: n1 => [|n] in eA; first by rewrite trivg_exponent eA in ntA.
+have nA1X: X \subset 'N('Ohm_1(A)) := char_norm_trans (Ohm_char 1 A) nAX.
+have sAnA1: 'Mho^n(A) \subset 'Ohm_1(A).
+ rewrite (MhoE n pA) (OhmE 1 pA) genS //.
+ apply/subsetP=> xpn; case/imsetP=> x Ax ->{xpn}; rewrite !inE groupX //.
+ by rewrite -expgM -expnSr -eA -order_dvdn dvdn_exponent.
+have nAnX: X \subset 'N('Mho^n(A)) := char_norm_trans (Mho_char n A) nAX.
+have [B minB sBAn]: {B : {group gT} | minnormal B X & B \subset 'Mho^n(A)}.
+ apply: mingroup_exists; rewrite nAnX andbT; apply/trivgPn.
+ have [x Ax ox] := exponent_witness (abelian_nil cAA).
+ exists (x ^+ (p ^ n)); first by rewrite Mho_p_elt ?(mem_p_elt pA).
+ by rewrite -order_dvdn -ox eA dvdn_Pexp2l ?ltnn.
+have abelA1: p.-abelem 'Ohm_1(A) by rewrite Ohm1_abelem.
+have sBA1: B \subset 'Ohm_1(A) := subset_trans sBAn sAnA1.
+case/mingroupP: minB; case/andP=> ntB nBX minB.
+have{nBX sBA1} [U defA1 nUX] := Maschke_abelem abelA1 p'X sBA1 nA1X nBX.
+have [_ mulBU _ tiBU] := dprodP defA1; have{mulBU} [_ sUA1] := mulG_sub mulBU.
+have sUA: U \subset A := subset_trans sUA1 (Ohm_sub 1 _).
+have [U1 | {defA1 minB}ntU] := eqsVneq U 1.
+ rewrite U1 dprodg1 /= in defA1.
+ have homoA: homocyclic A.
+ apply/(Ohm1_homocyclicP pA cAA); rewrite eA pfactorK //=.
+ by apply/eqP; rewrite eqEsubset sAnA1 -defA1 sBAn.
+ exists [set A]; rewrite ?big_set1 // => G; move/set1P->; split=> //.
+ have OhmMho: forall k, 'Ohm_k(A) = 'Mho^(n.+1 - k)(A).
+ by move=> k; rewrite (homocyclic_Ohm_Mho k pA) // eA pfactorK.
+ have fM: {in A &, {morph (fun x => x ^+ (p ^ n)) : x y / x * y >-> x * y}}.
+ by move=> x y Ax Ay /=; rewrite expgMn // /commute -(centsP cAA).
+ pose f := Morphism fM; have ker_f: 'ker f = 'Phi(A).
+ apply/setP=> z; rewrite (Phi_Mho pA cAA) -(subSnn n) -OhmMho.
+ by rewrite (OhmEabelian pA) ?(abelianS (Ohm_sub n A)) ?inE.
+ have [g injg def_g] := first_isom f; rewrite /= {}ker_f in g injg def_g.
+ have{f def_g} def_g: forall H, gval H \subset A -> g @* (H / _) = 'Mho^n(H).
+ move=> H sHA; rewrite def_g morphimEsub //.
+ by rewrite (MhoEabelian n (pgroupS sHA pA) (abelianS sHA cAA)).
+ have im_g: g @* (A / 'Phi(A)) = B by rewrite def_g // defA1 OhmMho subn1.
+ have defAb: A / 'Phi(A) = g @*^-1 B by rewrite -im_g injmK.
+ have nsPhiA: 'Phi(A) <| A := Phi_normal A.
+ have nPhiX: X \subset 'N('Phi(A)) := char_norm_trans (Phi_char A) nAX.
+ rewrite defAb; apply/mingroupP; split=> [|Hb].
+ by rewrite -(morphim_injm_eq1 injg) ?morphpreK /= -?defAb ?im_g ?ntB ?actsQ.
+ case/andP=> ntHb actsXHb /= sgHbB; have [sHbA _] := subsetIP sgHbB.
+ rewrite -sub_morphim_pre // in sgHbB; rewrite -(minB _ _ sgHbB) ?injmK //.
+ rewrite morphim_injm_eq1 // {}ntHb {actsXHb}(subset_trans actsXHb) //=.
+ have{sHbA} [H defHb sPhiH sHA] := inv_quotientS nsPhiA sHbA.
+ rewrite defHb def_g // (char_norm_trans (Mho_char n H)) //.
+ by rewrite astabsQ ?subsetIr ?(normalS sPhiH sHA).
+have nsUA: U <| A by rewrite -sub_abelian_normal.
+have nUA: A \subset 'N(U) by case/andP: nsUA.
+have Au_lt_m: #|A / U| < m := leq_trans (ltn_quotient ntU sUA) leAm.
+have cAuAu: abelian (A / U) := quotient_abelian _ cAA.
+have pAu: p.-group (A / U) := quotient_pgroup _ pA.
+have p'Xu: p^'.-group (X / U) := quotient_pgroup _ p'X.
+have nXAu: X / U \subset 'N(A / U) := quotient_norms _ nAX.
+have{Au_lt_m p'Xu nXAu} [S defAu simS] := IHm _ _ _ Au_lt_m cAuAu pAu p'Xu nXAu.
+have sSAu: forall Ku, Ku \in S -> Ku \subset A / U.
+ by move=> Ku S_Ku; rewrite -(bigdprodWY defAu) sub_gen // (bigcup_max Ku).
+have{B ntB sBAn tiBU} [Ku S_Ku eKu]: exists2 Ku, Ku \in S & exponent Ku == (p ^ n.+1)%N.
+ apply/exists_inP; apply: contraR ntB; rewrite negb_exists_in -subG1 -tiBU.
+ move/forall_inP=> expSpn; apply/subsetP=> x Ux; rewrite inE Ux coset_idr //.
+ by rewrite (subsetP nUA) // (subsetP (Mho_sub n A)) // (subsetP sBAn).
+ have [y Ay ->]: exists2 y, y \in A & x = y ^+ (p ^ n).
+ by apply/imsetP; rewrite -MhoEabelian ?(subsetP sBAn).
+ rewrite morphX ?(subsetP nUA) // (exponentP _ _ (mem_quotient _ Ay)) //.
+ rewrite -sub_Ldiv -OhmEabelian ?(abelianS (Ohm_sub n _)) //=.
+ rewrite (OhmE n pAu) /= -(bigdprodWY defAu) genS // subsetI sub_gen //=.
+ apply/bigcupsP=> Ku S_Ku; rewrite sub_LdivT.
+ have: exponent Ku %| p ^ n.+1.
+ by rewrite (dvdn_trans (exponentS (sSAu _ S_Ku))) // -eA exponent_quotient.
+ case/dvdn_pfactor=> // k le_k_n1 expKu; rewrite expKu dvdn_exp2l //.
+ by rewrite -ltnS ltn_neqAle le_k_n1 -(eqn_exp2l _ _ p_gt1) -expKu expSpn.
+have{sSAu} [sKuA [homoKu nKuX minKu]] := (sSAu Ku S_Ku, simS Ku S_Ku).
+have [K defKu sUK sKA] := inv_quotientS nsUA sKuA.
+have [cKK cKuKu] := (abelianS sKA cAA, abelianS sKuA cAuAu).
+have [pK pKu] := (pgroupS sKA pA, pgroupS sKuA pAu).
+have nsUK: U <| K := normalS sUK sKA nsUA; have [_ nUK] := andP nsUK.
+have nKX: X \subset 'N(K).
+ by rewrite -(quotientSGK nUX) ?normsG ?quotient_normG // -defKu.
+pose K1 := 'Mho^1(K); have nsK1K: K1 <| K := Mho_normal 1 K.
+have nXKb: X / K1 \subset 'N(K / K1) by exact: quotient_norms.
+pose K'u := \big[dprod/1]_(Bu in S :\ Ku) Bu.
+have{S_Ku} defAu_K: K / U \x K'u = A / U by rewrite -defKu -big_setD1.
+have [[_ Pu _ defK'u]] := dprodP defAu_K; rewrite defK'u => mulKPu _ tiKPu.
+have [_ sPuA] := mulG_sub mulKPu.
+have [P defPu sUP sPA] := inv_quotientS nsUA sPuA.
+have{simS defK'u} nPX: X \subset 'N(P).
+ rewrite -(quotientSGK nUX) ?normsG // quotient_normG ?(normalS sUP sPA) //.
+ rewrite -defPu -(bigdprodWY defK'u) norms_gen ?norms_bigcup //.
+ by apply/bigcapsP=> Bu; case/setD1P=> _; case/simS.
+have abelKb: p.-abelem (K / K1).
+ by rewrite -[K1](Phi_Mho pK) ?Phi_quotient_abelem.
+have p'Xb: p^'.-group (X / K1) := quotient_pgroup _ p'X.
+have sUKb: U / K1 \subset K / K1 := quotientS _ sUK.
+have nUXb: X / K1 \subset 'N(U / K1) := quotient_norms _ nUX.
+have tiUK1: U :&: K1 = 1.
+ apply/trivgP; apply/subsetP=> xp; case/setIP=> Uxp K1xp.
+ have{K1xp} [x Kx def_xp]: exists2 x, x \in K & xp = x ^+ p.
+ by apply/imsetP; rewrite -(MhoEabelian 1).
+ suffices A1x: x \in 'Ohm_1(A).
+ by rewrite def_xp inE; case/abelemP: abelA1 => // _ ->.
+ have nUx: x \in 'N(U) := subsetP nUK x Kx.
+ rewrite -sub1set -(quotientSGK _ sUA1) ?quotient_set1 ?sub1set //.
+ apply: (subsetP (quotientS U (subset_trans (MhoS n sKA) sAnA1))).
+ rewrite quotientE morphim_Mho //= -quotientE -defKu.
+ have ->: 'Mho^n(Ku) = 'Ohm_1(Ku).
+ by rewrite (homocyclic_Ohm_Mho 1 pKu) // (eqP eKu) pfactorK ?subn1.
+ rewrite (OhmE 1 pKu) ?mem_gen // !inE defKu mem_quotient //=.
+ by rewrite -morphX //= -def_xp coset_id.
+have [Db defKb nDXb] := Maschke_abelem abelKb p'Xb sUKb nXKb nUXb.
+have [_ mulUDb _ tiUDb] := dprodP defKb; have [_ sDKb] := mulG_sub mulUDb.
+have [D defDb sK1D sDK] := inv_quotientS (Mho_normal 1 K) sDKb.
+have nK1X: X \subset 'N(K1) := char_norm_trans (Mho_char 1 K) nKX.
+have [cDU [sK1K nK1K]] := (centSS sUK sDK cKK, andP nsK1K).
+have nDX: X \subset 'N(D).
+ rewrite -(quotientSGK nK1X) ?normsG // quotient_normG ?(normalS _ sDK) //.
+ by rewrite -defDb.
+have{mulUDb} mulUD: U * D = K.
+ rewrite (centC cDU) -(mulSGid sK1D) -mulgA -(centC cDU).
+ rewrite -quotientK ?quotientMr ?(subset_trans _ nK1K) ?mul_subG // -defDb.
+ by rewrite mulUDb quotientGK.
+have cKP: P \subset 'C(K) := centSS sPA sKA cAA.
+have mulKP: K * P = A.
+ rewrite -(mulSGid sUK) -mulgA -(quotientGK nsUA) -mulKPu defPu.
+ by rewrite -quotientK ?quotientMr ?mul_subG ?(subset_trans _ nUA).
+have defKP: K :&: P = U.
+ apply/eqP; rewrite eqEsubset subsetI sUK sUP !andbT.
+ by rewrite -quotient_sub1 ?subIset ?nUK // -tiKPu defPu quotientI.
+have tiUD: U :&: D = 1.
+ apply/trivgP; rewrite -tiUK1 subsetI subsetIl.
+ rewrite -quotient_sub1; last by rewrite subIset ?(subset_trans sUK).
+ by rewrite -tiUDb defDb quotientI.
+have tiDP: D :&: P = 1 by rewrite -(setIidPl sDK) -setIA defKP setIC.
+have mulDP: D * P = A by rewrite -(mulSGid sUP) mulgA -(centC cDU) mulUD.
+have sDA := subset_trans sDK sKA.
+have defA: D \x P = A by rewrite dprodE // (centSS sPA sDA).
+have ntD: D :!=: 1.
+ apply: contraNneq ntA => D1; rewrite trivg_exponent eA -(eqP eKu).
+ rewrite -trivg_exponent -subG1 -tiKPu defKu subsetIidl defPu quotientS //.
+ by rewrite -(mul1g P) -D1 mulDP.
+have ltPm: #|P| < m.
+ by rewrite (leq_trans _ leAm) // -(dprod_card defA) ltn_Pmull ?cardG_gt1.
+have [cPP pP] := (abelianS sPA cAA, pgroupS sPA pA).
+have{S defAu K'u defAu_K} [S defP simS] := IHm _ _ _ ltPm cPP pP p'X nPX.
+exists (D |: S) => [ | {defP}B].
+ rewrite big_setU1 ?defP //=; apply: contra ntD => S_D.
+ by rewrite -subG1 -tiDP subsetIidl -(bigdprodWY defP) sub_gen ?(bigcup_max D).
+case/setU1P=> [-> {B S simS} | ]; last exact: simS.
+have [[pD cDD] nUD] := (pgroupS sDA pA, abelianS sDA cAA, subset_trans sDA nUA).
+have isoD: D \isog Ku by rewrite defKu -mulUD quotientMidl quotient_isog.
+rewrite {isoD}(isog_homocyclic isoD); split=> //.
+have nPhiDX: X \subset 'N('Phi(D)) := char_norm_trans (Phi_char D) nDX.
+have [f [injf im_f act_f]]:
+ exists f : {morphism D / 'Phi(D) >-> coset_of 'Phi(Ku)},
+ [/\ 'injm f, f @* (D / 'Phi(D)) = Ku / 'Phi(Ku)
+ & {in D / 'Phi(D) & X, morph_act 'Q 'Q f (coset U)}].
+- have [/= injf im_f] := isomP (quotient_isom nUD tiUD).
+ set f := restrm nUD (coset U) in injf im_f.
+ rewrite -quotientMidl mulUD -defKu in im_f.
+ have fPhiD: f @* 'Phi(D) = 'Phi(Ku) by rewrite -im_f (morphim_Phi _ pD).
+ rewrite -['Phi(Ku)]/(gval 'Phi(Ku)%G) -(group_inj fPhiD).
+ exists (quotm_morphism [morphism of f] (Phi_normal _)).
+ rewrite (injm_quotm _ injf) morphim_quotm /= -/f im_f.
+ split=> // yb x; case/morphimP=> y Ny Dy ->{yb} Xx.
+ have [nPhiDx nUx] := (subsetP nPhiDX x Xx, subsetP nUX x Xx).
+ have Dyx: y ^ x \in D by rewrite memJ_norm // (subsetP nDX).
+ rewrite quotmE // !qactE ?qact_domE ?subsetT ?astabsJ ?quotmE //=.
+ - by congr (coset _ _); rewrite /f /restrm morphJ // (subsetP nUD).
+ - by rewrite (subsetP (morphim_norm _ _)) ?mem_morphim.
+ rewrite morphim_restrm (setIidPr (Phi_sub _)).
+ by rewrite (subsetP (morphim_norm _ _)) ?mem_quotient.
+apply/mingroupP; split=> [|Y].
+ rewrite -subG1 quotient_sub1 ?(normal_norm (Phi_normal _)) //.
+ by rewrite proper_subn ?Phi_proper // actsQ.
+case/andP=> ntY actsXY sYD; have{minKu} [_ minKu] := mingroupP minKu.
+apply: (injm_morphim_inj injf); rewrite // im_f.
+apply: minKu; last by rewrite /= -im_f morphimS.
+rewrite morphim_injm_eq1 // ntY.
+apply/subsetP=> xb; case/morphimP=> x Nx Xx ->{xb}.
+rewrite 2!inE /= qact_domE ?subsetT // astabsJ.
+rewrite (subsetP (char_norm_trans (Phi_char _) nKuX)) ?mem_quotient //=.
+apply/subsetP=> fy; case/morphimP=> y Dy Yy ->{fy}.
+by rewrite inE /= -act_f // morphimEsub // mem_imset // (acts_act actsXY).
+Qed.
+
+CoInductive is_iso_quotient_homocyclic_sdprod gT (V G : {group gT}) m : Prop :=
+ IsoQuotientHomocyclicSdprod wT (W D G1 : {group wT}) (f : {morphism D >-> gT})
+ of homocyclic W & #|W| = (#|V| ^ m)%N
+ & 'ker f = 'Mho^1(W) & f @* W = V & f @* G1 = G & W ><| G1 = D.
+
+Lemma iso_quotient_homocyclic_sdprod gT (V G : {group gT}) p m :
+ minnormal V G -> coprime p #|G| -> p.-abelem V -> m > 0 ->
+ is_iso_quotient_homocyclic_sdprod V G m.
+Proof.
+move=> minV copG abelV m_gt0; pose q := (p ^ m)%N.
+have [ntV nVG] := andP (mingroupp minV).
+have [p_pr pVdvdn [n Vpexpn]] := pgroup_pdiv (abelem_pgroup abelV) ntV.
+move/(abelem_mx_irrP abelV ntV nVG): (minV) => mx_irrV.
+have dim_lt0 : 'dim V > 0 by rewrite (dim_abelemE abelV) // Vpexpn pfactorK.
+have q_gt1: q > 1 by rewrite (ltn_exp2l 0) // prime_gt1.
+have p_q: p.-nat q by rewrite pnat_exp pnat_id.
+have p_dv_q: p %| q := dvdn_exp2l p m_gt0.
+pose rG := regular_repr [comUnitRingType of 'Z_q] G; pose MR_G := ('MR rG)%gact.
+have [wT [fL injL [fX injX fJ]]]: exists wT : finGroupType,
+ exists2 fL : {morphism setT >-> wT}, 'injm fL &
+ exists2 fX : {morphism G >-> wT}, 'injm fX &
+ {in setT & G, morph_act MR_G 'J fL fX}.
+- exists (sdprod_groupType MR_G).
+ exists (sdpair1_morphism MR_G); first exact: injm_sdpair1.
+ by exists (sdpair2_morphism MR_G); [exact: injm_sdpair2 | exact: sdpair_act].
+move imfL: (fL @* [set: _])%G => L; move imfX: (fX @* G)%G => X.
+have cLL: abelian L by rewrite -imfL morphim_abelian // zmod_abelian.
+have pL: p.-group L.
+ by rewrite -imfL morphim_pgroup -?pnat_exponent ?exponent_mx_group.
+have tiVG: V :&: G = 1 by rewrite coprime_TIg // Vpexpn coprime_pexpl.
+have{copG} p'G: p^'.-group G by rewrite /pgroup p'natE // -prime_coprime.
+have p'X: p^'.-group X by rewrite -imfX morphim_pgroup.
+have nXL: X \subset 'N(L).
+ rewrite -imfX -imfL; apply/subsetP=> _ /morphimP[x Gx _ ->]; rewrite inE.
+ apply/subsetP=> _ /imsetP[_ /morphimP[v rVv _ ->] ->].
+ by rewrite -fJ // mem_morphim ?in_setT.
+have [/= S defL im_S] := coprime_act_abelian_pgroup_structure cLL pL p'X nXL.
+pose gi (z : 'Z_q) := z%:R : 'F_p.
+have giM: rmorphism gi.
+ split=> [z1 z2|]; last split=> // z1 z2.
+ apply: canRL (addrK _) _; apply: val_inj.
+ by rewrite -{2}(subrK z2 z1) -natrD /= !val_Fp_nat ?modn_dvdm // Zp_cast.
+ by apply: val_inj; rewrite -natrM /= !val_Fp_nat ?modn_dvdm // Zp_cast.
+have [gL [DgL _ _ _]] := domP (invm_morphism injL) (congr_group imfL).
+pose g u := map_mx (RMorphism giM) (gL u).
+have gM: {in L &, {morph g : u v / u * v}}.
+ by move=> u v Lu Lv /=; rewrite {1}/g morphM // map_mxD.
+have kerg: 'ker (Morphism gM) = 'Phi(L).
+ rewrite (Phi_Mho pL cLL) (MhoEabelian 1 pL cLL).
+ apply/setP=> u; apply/idP/imsetP=> [ | [v Lv ->{u}]]; last first.
+ rewrite !inE groupX //=; apply/eqP/rowP=> i; apply: val_inj.
+ rewrite !mxE morphX // mulmxnE Zp_mulrn /= val_Fp_nat //=.
+ by move: {i}(_ i); rewrite Zp_cast // => vi; rewrite modn_dvdm // modnMl.
+ case/morphpreP; rewrite -{1}imfL => /morphimP[v rVv _ ->{u}] /set1P /=.
+ rewrite /g DgL /= invmE // => /rowP vp0.
+ pose x := fL (map_mx (fun t : 'Z_q => (t %/ p)%:R) v).
+ exists x; first by rewrite -imfL mem_morphim ?inE.
+ rewrite -morphX ?in_setT //; congr (fL _); apply/rowP=> i.
+ rewrite mulmxnE -{1}(natr_Zp (v 0 i)) {1}(divn_eq (v 0 i) p) addnC.
+ by have:= congr1 val (vp0 i); rewrite !mxE -mulrnA /= val_Fp_nat // => ->.
+have [gX [DgX KgX _ imgX]] := domP (invm_morphism injX) (congr_group imfX).
+pose aG := regular_repr [fieldType of 'F_p] G.
+have GgX: {in X, forall x, gX x \in G}.
+ by rewrite DgX -imfX => _ /morphimP[x Gx _ ->]; rewrite /= invmE.
+have XfX: {in G, forall x, fX x \in X}.
+ by move=> x Gx; rewrite -imfX mem_morphim.
+have gJ: {in L & X, forall u x, g (u ^ x) = g u *m aG (gX x)}.
+ rewrite -{1}imfL -{1}imfX => _ _ /morphimP[u rVu _ ->] /morphimP[x Gx _ ->].
+ rewrite -fJ // /g DgL DgX /= !invmE // mx_repr_actE ?inE //.
+ by rewrite (map_mxM (RMorphism giM)) map_regular_mx.
+pose gMx U := rowg_mx (Morphism gM @* U).
+have simS: forall U, U \in S -> mxsimple aG (gMx U).
+ move=> U S_U; have [_ nUX irrU] := im_S U S_U.
+ have{irrU} [modU irrU] := mingroupP irrU; have{modU} [ntU _] := andP modU.
+ have sUL: U \subset L by rewrite -(bigdprodWY defL) sub_gen // (bigcup_max U).
+ split=> [||U2 modU2].
+ - rewrite (eqmx_module _ (genmxE _)); apply/mxmoduleP=> x Gx.
+ apply/row_subP=> i; rewrite row_mul rowK.
+ have [u Lu Uu def_u] := morphimP (enum_valP i).
+ rewrite -(invmE injX Gx) -DgX def_u -gJ ?XfX //.
+ set ux := u ^ _; apply: eq_row_sub (gring_index _ (g ux)) _.
+ have Uux: ux \in U by rewrite memJ_norm // (subsetP nUX) ?XfX.
+ by rewrite rowK gring_indexK //; apply: mem_morphim; rewrite ?(subsetP sUL).
+ - apply: contra ntU; rewrite rowg_mx_eq0.
+ rewrite -subG1 sub_morphim_pre // -kerE kerg => sU_Phi.
+ rewrite /= quotientS1 //=; rewrite (big_setD1 U) //= in defL.
+ have{defL} [[_ U' _ ->] defUU' cUU' tiUU'] := dprodP defL.
+ have defL: U \* U' = L by rewrite cprodE.
+ have:= cprod_modl (Phi_cprod pL defL) (Phi_sub U).
+ rewrite -(setIidPl (Phi_sub U')) setIAC -setIA tiUU' setIg1 cprodg1 => ->.
+ by rewrite subsetIidr.
+ rewrite -!rowgS stable_rowg_mxK /= => [sU2gU nzU2|]; last first.
+ apply/subsetP=> z _; rewrite !inE /=; apply/subsetP=> u gUu.
+ by rewrite inE /= /scale_act -[val z]natr_Zp scaler_nat groupX.
+ rewrite sub_morphim_pre // -subsetIidl.
+ rewrite -(quotientSGK (normal_norm (Phi_normal U))) //=; last first.
+ rewrite subsetI Phi_sub (subset_trans (PhiS pL sUL)) //.
+ by rewrite -kerg ker_sub_pre.
+ rewrite [(U :&: _) / _]irrU //; last by rewrite quotientS ?subsetIl.
+ rewrite (contra _ nzU2) /=; last first.
+ rewrite -subG1 quotient_sub1; last first.
+ by rewrite subIset // normal_norm // Phi_normal.
+ rewrite /= -(morphpre_restrm sUL).
+ move/(morphimS (restrm_morphism sUL (Morphism gM))).
+ rewrite morphpreK ?im_restrm // morphim_restrm => s_U2_1.
+ rewrite -trivg_rowg -subG1 (subset_trans s_U2_1) //.
+ rewrite -(morphim_ker (Morphism gM)) morphimS // kerg.
+ by rewrite subIset ?(PhiS pL) ?orbT.
+ rewrite actsQ //; first by rewrite (char_norm_trans (Phi_char U)).
+ rewrite normsI //; apply/subsetP=> x Xx; rewrite inE.
+ apply/subsetP=> _ /imsetP[u g'U2u ->].
+ have [Lu U2gu] := morphpreP g'U2u; rewrite mem_rowg in U2gu.
+ rewrite inE memJ_norm ?(subsetP nXL) // Lu /= inE gJ //.
+ by rewrite mem_rowg (mxmodule_trans modU2) ?GgX.
+have im_g: Morphism gM @* L = [set: 'rV_#|G|].
+ apply/eqP; rewrite eqEcard subsetT cardsT card_matrix card_Fp //= mul1n.
+ rewrite card_morphim kerg setIid (Phi_Mho pL cLL) -divgS ?Mho_sub //.
+ rewrite -(mul_card_Ohm_Mho_abelian 1 cLL) mulnK ?cardG_gt0 //.
+ rewrite (card_pgroup (pgroupS (Ohm_sub 1 L) pL)) -rank_abelian_pgroup //.
+ by rewrite -imfL (injm_rank injL) //= rank_mx_group mul1n.
+have sumS: (\sum_(U in S) gMx U :=: 1%:M)%MS.
+ apply/eqmxP; rewrite submx1; apply/rV_subP=> v _.
+ have: v \in Morphism gM @* L by rewrite im_g inE.
+ case/morphimP=> u Lu _ ->{v}.
+ rewrite -mem_rowg -sub1set -morphim_set1 // sub_morphim_pre ?sub1set //.
+ have [c [Uc -> _]] := mem_bigdprod defL Lu; rewrite group_prod //= => U S_U.
+ have sUL: U \subset L by rewrite -(bigdprodWY defL) sub_gen // (bigcup_max U).
+ rewrite inE (subsetP sUL) ?Uc // inE mem_rowg (sumsmx_sup U) // -mem_rowg.
+ by rewrite (subsetP (sub_rowg_mx _)) // mem_morphim ?(subsetP sUL) ?Uc.
+have Fp'G: [char 'F_p]^'.-group G.
+ by rewrite (eq_p'group _ (charf_eq (char_Fp p_pr))).
+have [VK [modVK defVK]] := rsim_regular_submod mx_irrV Fp'G.
+have [U S_U isoUV]: {U | U \in S & mx_iso (regular_repr _ G) (gMx U) VK}.
+ apply: hom_mxsemisimple_iso (scalar_mx_hom _ 1 _) _ => [|U S_U _|]; auto.
+ by apply/(submod_mx_irr modVK); exact: (mx_rsim_irr defVK).
+ by rewrite mulmx1 sumS submx1.
+have simU := simS U S_U; have [modU _ _] := simU.
+pose rV := abelem_repr abelV ntV nVG.
+have{VK modVK defVK isoUV} [h dimU h_free hJ]: mx_rsim (submod_repr modU) rV.
+ by apply: mx_rsim_trans (mx_rsim_sym defVK); exact/mx_rsim_iso.
+have sUL : U \subset L.
+ by move: defL; rewrite (big_setD1 U) //= => /dprodP[[_ U1 _ ->] /mulG_sub[]].
+pose W := [set: 'rV['Z_(p ^ m)](V)]%G.
+have [homU nUX _] := im_S U S_U; have [cUU _] := andP homU.
+have atypeU: abelian_type U == nseq ('dim V) (p ^ m)%N.
+ rewrite (big_setD1 U) //= in defL.
+ have [[_ U' _ defU'] defUU' _ tiUU'] := dprodP defL.
+ rewrite defU' in defL defUU' tiUU'.
+ have ->: 'dim V = 'r(U).
+ apply/eqP; rewrite -dimU -(eqn_exp2l _ _ (prime_gt1 p_pr)).
+ rewrite (rank_abelian_pgroup (pgroupS sUL pL) cUU).
+ rewrite -(card_pgroup (pgroupS (Ohm_sub 1 U) (pgroupS sUL pL))).
+ rewrite -{1}(card_Fp p_pr) -card_rowg stable_rowg_mxK; last first.
+ apply/subsetP=> z _; rewrite !inE; apply/subsetP=> v gUv.
+ by rewrite inE /= /scale_act -(natr_Zp (val z)) scaler_nat groupX.
+ rewrite card_morphim kerg (Phi_Mho pL cLL) (setIidPr sUL) -divgI setIC.
+ rewrite -(dprod_modl (Mho_dprod 1 defL) (Mho_sub 1 U)).
+ rewrite [_ :&: _](trivgP _); last by rewrite -tiUU' setIC setSI ?Mho_sub.
+ by rewrite dprodg1 -(mul_card_Ohm_Mho_abelian 1 cUU) mulnK ?cardG_gt0.
+ have isoL: isog L [set: 'rV['Z_q]_#|G|] by rewrite -imfL isog_sym sub_isog.
+ have homL: homocyclic L by rewrite (isog_homocyclic isoL) mx_group_homocyclic.
+ have [-> _] := abelian_type_dprod_homocyclic defL pL homL.
+ by rewrite (exponent_isog isoL) // exponent_mx_group.
+have pU: p.-group U by rewrite (pgroupS sUL).
+have isoWU: isog U W.
+ by rewrite eq_abelian_type_isog ?zmod_abelian // abelian_type_mx_group ?mul1n.
+have {cUU atypeU} cardU : #|U| = (#|V| ^ m)%N.
+ rewrite card_homocyclic // (exponent_isog isoWU) exponent_mx_group //.
+ rewrite -size_abelian_type // (eqP atypeU) size_nseq.
+ by rewrite (dim_abelemE abelV) // Vpexpn pfactorK // expnAC.
+pose f3 w := rVabelem abelV ntV (in_submod _ (g w) *m h).
+have f3M: {in U &, {morph f3: w1 w2 / w1 * w2}}.
+ move=> w1 w2 Uw1 Uw2 /=; rewrite {1}/f3.
+ rewrite gM ?(subsetP sUL) // linearD mulmxDl.
+ by rewrite morphM ?mem_im_abelem_rV.
+have Ugw w : w \in U -> (g w <= gMx U)%MS.
+ move=> Uw; rewrite -mem_rowg (subsetP (sub_rowg_mx _)) //.
+ by rewrite (mem_morphim (Morphism gM)) ?(subsetP sUL).
+have KgU: 'ker_U (Morphism gM) = 'Mho^1(U).
+ rewrite kerg (Phi_Mho pL cLL); rewrite (big_setD1 U) //= in defL.
+ have [[_ U' _ defU'] _ _ tiUU'] := dprodP defL; rewrite defU' in defL tiUU'.
+ rewrite setIC -(dprod_modl (Mho_dprod 1 defL) (Mho_sub 1 U)).
+ by rewrite [_ :&: _](trivgP _) ?dprodg1 // -tiUU' setIC setSI ?Mho_sub.
+have{KgU} Kf3: 'ker (Morphism f3M) = 'Mho^1(U).
+ apply/setP=> w; rewrite !inE /=.
+ rewrite morph_injm_eq1 ?rVabelem_injm ?mem_im_abelem_rV //=.
+ rewrite -[1](mul0mx _ h) (inj_eq (row_free_inj h_free)) in_submod_eq0.
+ case Uw: (w \in U) => /=; last first.
+ by apply/sym_eq; apply: contraFF Uw; apply: (subsetP (Mho_sub _ _)).
+ rewrite -[(_ <= _)%MS]andTb -(Ugw _ Uw) -sub_capmx capmx_compl submx0.
+ by rewrite -KgU !inE Uw (subsetP sUL).
+have cUU: abelian U := abelianS sUL cLL.
+have im_f3: Morphism f3M @* U = V.
+ apply/eqP; rewrite eqEcard card_morphim setIid Kf3; apply/andP; split.
+ by apply/subsetP=> _ /morphimP[w _ _ ->]; apply: mem_rVabelem.
+ rewrite -divgS ?Mho_sub // -(mul_card_Ohm_Mho_abelian 1 cUU).
+ rewrite mulnK ?cardG_gt0 // (card_pgroup (pgroupS (Ohm_sub 1 U) pU)).
+ rewrite -rank_abelian_pgroup // (isog_rank isoWU) /W.
+ by rewrite (dim_abelemE abelV) // rank_mx_group mul1n Vpexpn pfactorK.
+have f3J: {in U & X, morph_act 'J 'J (Morphism f3M) gX}.
+ move=> u x Uu Xx; rewrite /f3 /= gJ ?(subsetP sUL) // in_submodJ ?Ugw //.
+ by rewrite -mulmxA hJ ?GgX // mulmxA rVabelemJ ?GgX.
+have defUX: U ><| X = U <*> X.
+ rewrite norm_joinEr; last by case: (im_S _ S_U).
+ by rewrite sdprodE ?coprime_TIg ?(pnat_coprime pU).
+pose f := sdprodm defUX f3J.
+have{im_f3} fU_V: f @* U = V by rewrite morphim_sdprodml.
+have fX_G: f @* X = G by rewrite morphim_sdprodmr // imgX -imfX im_invm.
+suffices: 'ker f = 'Mho^1(U) by exists wT U (U <*> X)%G X [morphism of f].
+rewrite -Kf3; apply/setP=> y; apply/idP/idP; last first.
+ move=> /morphpreP[/= Uy /set1P f3y].
+ by rewrite !inE /= sdprodmEl //= f3y (subsetP (joing_subl _ X)) /=.
+rewrite ker_sdprodm => /imset2P[u t Uu /setIdP[Xt /eqP/= fu] ->{y}].
+have: f3 u \in V :&: G.
+ by rewrite inE -fU_V morphim_sdprodml //= mem_imset ?setIid // fu GgX.
+rewrite tiVG in_set1 fu morph_injm_eq1 ?KgX ?injm_invm // => /eqP t1.
+by rewrite t1 invg1 mulg1 !inE Uu /= fu t1 morph1.
+Qed.
+
+Theorem solvable_Wielandt_fixpoint (I : finType) gT (A : I -> {group gT})
+ (n m : I -> nat) (G V : {group gT}) :
+ (forall i, m i + n i > 0 -> A i \subset G) ->
+ G \subset 'N(V) -> coprime #|V| #|G| -> solvable V ->
+ {in G, forall a, \sum_(i | a \in A i) m i = \sum_(i | a \in A i) n i}%N ->
+ (\prod_i #|'C_V(A i)| ^ (m i * #|A i|)
+ = \prod_i #|'C_V(A i)| ^ (n i * #|A i|))%N.
+Proof.
+move: {2}_.+1 (ltnSn #|V|) => c leVc sA_G nVG coVG solV partG; move: leVc.
+pose nz_k i := (0 < m i + n i)%N; rewrite !(bigID nz_k xpredT) /= {2 4}/nz_k.
+rewrite !(big1 _ (predC _)) /= => [|i|i]; try by case: (m i) (n i) => [[]|].
+pose sum_k A_ a k := (\sum_(i | (a \in (A_ i : {set _})) && nz_k i) k i)%N.
+have{partG} partG: {in G, forall a, sum_k _ A a m = sum_k _ A a n}.
+ move=> a /partG; rewrite !(bigID nz_k (fun i => a \in _)) -!/(sum_k _ A a _).
+ by rewrite /= !big1 ?addn0 /nz_k // => i /andP[_]; case: (m i) (n i) => [[]|].
+rewrite !muln1; elim: c => // c IHc in gT G A V nVG coVG solV partG sA_G *.
+rewrite ltnS => leVc; have [-> | ntV] := eqsVneq V 1.
+ by rewrite !big1 // => i _; rewrite setI1g cards1 exp1n.
+have nsVVG: V <| V <*> G by rewrite normalYl.
+without loss{c leVc IHc} minV: / minnormal V (V <*> G).
+ have [B minB sBV]: {B : {group gT} | minnormal B (V <*> G) & B \subset V}.
+ by apply: mingroup_exists; rewrite ntV normal_norm.
+ have [nBVG ntB abB] := minnormal_solvable minB sBV solV.
+ have [nBV nBG] := joing_subP nBVG; have solB := solvableS sBV solV.
+ have [{1}<- -> // | ltBV _] := eqVproper sBV.
+ have ltBc: #|B| < c := leq_trans (proper_card ltBV) leVc.
+ have coBG: coprime #|B| #|G| := coprimeSg sBV coVG.
+ have factorCA_B k i: nz_k i ->
+ (#|'C_B(A i)| ^ (k i * #|A i|) * #|'C_(V / B)(A i / B)| ^ (k i * #|A i / B|)
+ = #|'C_V(A i)| ^ (k i * #|A i|))%N.
+ - move/sA_G => sAiG; have nBAi := subset_trans sAiG nBG.
+ have [coVAi coBAi] := (coprimegS sAiG coVG, coprimegS sAiG coBG).
+ rewrite -(card_isog (quotient_isog _ _)) ?(coprime_TIg coBAi) // -expnMn.
+ rewrite -coprime_quotient_cent // -{1}(setIidPr sBV) setIAC.
+ by rewrite card_quotient ?LagrangeI // subIset ?nBV.
+ rewrite -!{1}(eq_bigr _ (factorCA_B _)) {factorCA_B} !big_split /=.
+ pose A_B i := A i / B; congr (_ * _)%N; first exact: (IHc _ G).
+ have: #|V / B| < c by apply: leq_trans leVc; rewrite ltn_quotient.
+ have (i): nz_k i -> A i / B \subset G / B by move/sA_G/quotientS->.
+ apply: IHc; rewrite ?morphim_sol ?coprime_morph ?quotient_norms //.
+ move=> _ /morphimP[a Na Ga ->].
+ suffices eqAB: sum_k _ A_B (coset B a) =1 sum_k _ A a by rewrite !eqAB partG.
+ move=> k; apply: eq_bigl => i; apply: andb_id2r => /sA_G sAiG.
+ rewrite -sub1set -quotient_set1 // quotientSK ?sub1set //.
+ by rewrite -{2}(mul1g (A i)) -(coprime_TIg coBG) setIC group_modr // inE Ga.
+have /is_abelemP[p p_pr abelV] := minnormal_solvable_abelem minV solV.
+have [p_gt1 [pV cVV _]] := (prime_gt1 p_pr, and3P abelV).
+have{minV} minV: minnormal V G.
+ apply/mingroupP; split=> [|B nBG sBV]; first by rewrite ntV nVG.
+ by case/mingroupP: minV => _ -> //; rewrite join_subG (sub_abelian_norm cVV).
+have co_pG: coprime p #|G|.
+ by have [_ _ [e oV]] := pgroup_pdiv pV ntV; rewrite oV coprime_pexpl in coVG.
+have p'G: p^'.-group G by rewrite pgroupE p'natE -?prime_coprime.
+pose rC i := logn p #|'C_V(A i)|.
+have ErC k i: (#|'C_V(A i)| ^ (k i * #|A i|) = p ^ (rC i * k i * #|A i|))%N.
+ suffices /card_pgroup->: p.-group 'C_V(A i) by rewrite -expnM mulnA.
+ by rewrite (pgroupS (subsetIl _ _)).
+rewrite !{1}(eq_bigr _ (fun i _ => ErC _ i)) {ErC} -!expn_sum; congr (_ ^ _)%N.
+have eqmodX x y: (forall e, x = y %[mod p ^ e]) -> x = y.
+ pose e := maxn x y; move/(_ e); have:= ltn_expl e p_gt1.
+ by rewrite gtn_max /= => /andP[x_ltq y_ltq]; rewrite !modn_small.
+apply: eqmodX => e; have [-> | e_gt0] := posnP e; first by rewrite !modn1.
+set q := (p ^ e)%N; have q_gt1: q > 1 by rewrite -(exp1n e) ltn_exp2r.
+have{e_gt0 co_pG} [wT W D G1 f homoW oW kerf imfW imfG1 defD] :=
+ iso_quotient_homocyclic_sdprod minV co_pG abelV e_gt0.
+have [[cWW _] [_ /mulG_sub[sWD sG1D] nWG1 tiWG1]] := (andP homoW, sdprodP defD).
+have pW: p.-group W by rewrite /pgroup oW pnat_exp [p.-nat _]pV.
+have rW_V: 'r(W) = 'dim V.
+ rewrite (rank_abelian_pgroup pW cWW) -(mulnK #|_| (cardG_gt0 'Mho^1(W))).
+ rewrite mul_card_Ohm_Mho_abelian // divg_normal ?Mho_normal //=.
+ rewrite -(setIidPr (Mho_sub 1 W)) -kerf.
+ by rewrite (card_isog (first_isog_loc _ _)) //= imfW (dim_abelemE abelV).
+have expW: exponent W = q.
+ apply/eqP; rewrite -(@eqn_exp2r _ _ ('dim V)) // -{1}rW_V -expnM mulnC expnM.
+ by rewrite (dim_abelemE abelV) -?card_pgroup // -oW eq_sym max_card_abelian.
+have{rW_V} /isogP[fW injfW im_fW]: [set: 'rV['Z_q](V)] \isog W.
+ rewrite eq_abelian_type_isog ?zmod_abelian // abelian_type_mx_group ?mul1n //.
+ by rewrite abelian_type_homocyclic // rW_V expW.
+have WfW u: fW u \in W by rewrite -im_fW mem_morphim ?inE.
+have [fW' [DfW' KfW' _ _]] := domP (invm_morphism injfW) im_fW.
+have{KfW'} injfW': 'injm fW' by rewrite KfW' injm_invm.
+have fW'K: {in W, cancel fW' fW} by move=> w Ww; rewrite DfW' invmK //= im_fW.
+have toWlin a1: linear (fun u => fW' (fW u ^ val (subg G1 a1))).
+ move=> z /= x y; rewrite (morphM fW) /= ?in_setT // conjMg /=.
+ rewrite morphM ?memJ_norm ?(subsetP nWG1) ?subgP //=; congr (_ * _).
+ rewrite -(natr_Zp z) !scaler_nat morphX ?in_setT // conjXg morphX //.
+ by rewrite memJ_norm // (subsetP nWG1) ?subgP.
+pose rW a1 := lin1_mx (Linear (toWlin a1)).
+pose fG := restrm sG1D f; have im_fG : fG @* G1 = G by rewrite im_restrm.
+have injfG: 'injm fG by rewrite -tiWG1 setIC ker_restrm kerf setIS ?Mho_sub.
+pose fG' := invm injfG; have im_fG': fG' @* G = G1 by rewrite -im_fG im_invm.
+pose gamma i := \sum_(a in A i) rW (fG' a).
+suffices{sum_k partG} tr_rW_Ai i: nz_k i -> \tr (gamma i) = (rC i * #|A i|)%:R.
+ have Dtr k i: nz_k i -> (rC i * k i * #|A i|)%:R = \tr (gamma i *+ k i).
+ by rewrite mulnAC natrM raddfMn mulr_natr /= => /tr_rW_Ai->.
+ rewrite -!val_Zp_nat // !natr_sum !{1}(eq_bigr _ (Dtr _)){Dtr}; congr (val _).
+ rewrite -!raddf_sum -!(eq_bigr _ (fun i _ => sumrMnl _ _ _ _)); congr (\tr _).
+ have sA_GP i a nz_i := subsetP (sA_G i nz_i) a.
+ rewrite !(exchange_big_dep (mem G)) {sA_GP}//=; apply: eq_bigr => a Ga.
+ by rewrite !sumrMnr !(big_andbC _ _ _ nz_k) -!/(sum_k _ A a _) partG.
+move/sA_G=> {sA_G} sAiG; pose Ai1 := fG' @* A i; pose rR := 'r([~: W, Ai1]).
+have sAiG1: Ai1 \subset G1 by rewrite -im_fG' morphimS.
+have AfG' a: a \in A i -> fG' a \in Ai1.
+ by move=> Aa; rewrite mem_morphim //= im_restrm imfG1 ?(subsetP sAiG).
+have coWAi1: coprime #|W| #|Ai1|.
+ by rewrite coprime_morphr ?(coprimegS sAiG) ?(pnat_coprime pW).
+suffices [Pl [Pr [Pu [Pd [PlrudK ErC ErR]]]]]:
+ exists Pl, exists Pr, exists Pu, exists Pd,
+ [/\ row_mx Pl Pr *m col_mx Pu Pd = 1%R,
+ {in A i, forall a, Pd *m (rW (fG' a) *m Pr) = 1%:M :> 'M_(rC i)}
+ & \sum_(a in A i) Pu *m (rW (fG' a) *m Pl) = 0 :> 'M_rR].
+- rewrite -(mulmx1 (gamma i)) idmxE -PlrudK mulmxA mxtrace_mulC mul_mx_row.
+ rewrite mul_col_row mxtrace_block !mulmx_suml !mulmx_sumr ErR mxtrace0 add0r.
+ by rewrite (eq_bigr _ ErC) sumr_const raddfMn /= mxtrace1 natrM mulr_natr.
+have defW: [~: W, Ai1] \x 'C_W(Ai1) = W.
+ by rewrite coprime_abelian_cent_dprod ?(subset_trans sAiG1).
+have [_ mulRCW _ tiRCW] := dprodP defW; have [sRW sCW] := mulG_sub mulRCW.
+have [homoRW homoCW] := dprod_homocyclic defW pW homoW.
+have [] := abelian_type_dprod_homocyclic defW pW homoW.
+rewrite expW -/rR => atypeRW atypeCW.
+have [[cRR _] [cCC _]] := (andP homoRW, andP homoCW).
+have{cRR atypeRW} /isogP[hR injhR im_hR]: [~: W, Ai1] \isog [set: 'rV['Z_q]_rR].
+ rewrite eq_abelian_type_isog ?zmod_abelian ?atypeRW //.
+ by rewrite abelian_type_mx_group // mul1n eqxx.
+have{tiRCW} rCW : 'r('C_W(Ai1)) = rC i.
+ rewrite -['r(_)]rank_Ohm1; have /rank_abelem ->: p.-abelem 'Ohm_1('C_W(Ai1)).
+ by rewrite Ohm1_abelem ?(pgroupS (subsetIl _ _)).
+ congr (logn p _); transitivity #|'C_W(Ai1) : 'Mho^1('C_W(Ai1))|.
+ by rewrite -divgS ?Mho_sub // -(mul_card_Ohm_Mho_abelian 1 cCC) mulnK.
+ transitivity #|'C_W(Ai1) : 'Mho^1(W)|.
+ symmetry; have /dprodP[_ /= defW1 _ _] := Mho_dprod 1 defW.
+ rewrite -indexgI; congr #|_ : _|; rewrite /= -defW1 -group_modr ?Mho_sub //.
+ by rewrite [_ :&: _](trivgP _) ?mul1g //= setIC -tiRCW setSI ?Mho_sub.
+ suffices /card_isog ->: 'C_V(A i) \isog 'C_W(Ai1) / 'Mho^1(W).
+ by rewrite card_quotient // subIset // normal_norm ?Mho_normal.
+ rewrite coprime_quotient_cent ?Mho_sub ?abelian_sol //= -/Ai1; last first.
+ by rewrite (subset_trans sAiG1) // (char_norm_trans _ nWG1) ?Mho_char.
+ have ->: A i :=: fG @* Ai1.
+ by rewrite /Ai1 morphim_invmE morphpreK // im_restrm imfG1.
+ rewrite -imfW morphim_restrm (setIidPr sAiG1).
+ have [f1 injf1 im_f1] := first_isom f.
+ rewrite -!im_f1 -injm_subcent ?quotientS ?(subset_trans sAiG1) //.
+ by rewrite -kerf isog_sym sub_isog // subIset ?quotientS.
+have{atypeCW} /isogP[hC injhC im_hC]: 'C_W(Ai1) \isog [set: 'rV['Z_q]_(rC i)].
+ rewrite eq_abelian_type_isog ?zmod_abelian // atypeCW rCW.
+ by rewrite abelian_type_mx_group ?mul1n.
+have mkMx m1 m2 (U : {group 'rV['Z_q]_m1}) (g : {morphism U >-> 'rV['Z_q]_m2}):
+ setT \subset 'dom g -> {Mg | mulmx^~ Mg =1 g}.
+- move/subsetP=> allU; suffices lin_g: linear g.
+ by exists (lin1_mx (Linear lin_g)) => u; rewrite mul_rV_lin1.
+ move=> z u v; rewrite morphM ?allU ?in_setT //.
+ by rewrite -(natr_Zp z) !scaler_nat -zmodXgE morphX ?allU ?in_setT.
+have /mkMx[Pu defPu]: setT \subset 'dom (invm injfW \o invm injhR).
+ by rewrite -sub_morphim_pre -im_hR // im_invm //= im_fW.
+have /mkMx[Pd defPd]: setT \subset 'dom (invm injfW \o invm injhC).
+ by rewrite -sub_morphim_pre -im_hC //= im_fW im_invm subsetIl.
+pose fUl := pairg1 [finGroupType of 'rV['Z_q]_(rC i)] \o hR.
+pose fUr := @pair1g [finGroupType of 'rV['Z_q]_rR] _ \o hC.
+have cRCW: fUr @* 'C_W(Ai1) \subset 'C(fUl @* [~: W, Ai1]).
+ rewrite !morphim_comp morphim_pair1g morphim_pairg1.
+ set UR := hR @* _; set UC := hC @* _.
+ by have/dprodP[] : _ = setX UR UC := setX_dprod _ _.
+have /domP[fUr' [DfUr' _ _ im_fUr']]: 'dom fUr = 'C_W(Ai1).
+ by rewrite /dom -im_hC injmK.
+have /domP[fUl' [DfUl' _ _ im_fUl']]: 'dom fUl = [~: W, Ai1].
+ by rewrite /dom -im_hR injmK.
+rewrite -{}im_fUr' -{}im_fUl' in cRCW; pose hW := dprodm defW cRCW.
+pose fPl := @fst _ _ \o (hW \o fW); pose fPr := @snd _ _ \o (hW \o fW).
+have /mkMx[/= Pl defPl]: setT \subset 'dom fPl.
+ by rewrite -!sub_morphim_pre ?subsetT ?im_fW.
+have /mkMx[/= Pr defPr]: setT \subset 'dom fPr.
+ by rewrite -!sub_morphim_pre ?subsetT ?im_fW.
+exists Pl, Pr, Pu, Pd; split.
+- apply/row_matrixP=> j; rewrite rowE -row1 mul_row_col mulmxDr !mulmxA.
+ apply: (injmP injfW); rewrite ?in_setT // morphM ?in_setT //.
+ rewrite defPl defPr defPu defPd -/hW [hW]lock /= -lock.
+ have /(mem_dprod defW)[jR [jC [RjR CjC -> _]]]:= WfW (row j 1).
+ rewrite [hW _]dprodmE // DfUl' DfUr' /= mulg1 mul1g !invmE // -DfW'.
+ by rewrite !fW'K ?(subsetP sRW jR) ?(subsetP sCW).
+- move=> a Aa; apply/row_matrixP=> j; pose jC := invm injhC (row j 1%:M).
+ rewrite rowE -row1 !mulmxA defPd defPr -/hW [hW]lock /= mul_rV_lin1 /= -lock.
+ have CjC: jC \in 'C_W(Ai1).
+ by rewrite -(im_invm injhC) mem_morphim /= ?im_hC ?inE.
+ have [[/fW'K id_jC /centP cA1jC] A1a] := (setIP CjC, AfG' a Aa).
+ rewrite -DfW' id_jC subgK ?(subsetP sAiG1) // /conjg cA1jC // mulKg id_jC.
+ by rewrite [hW _]dprodmEr ?DfUr' //= invmK ?im_hC ?inE.
+apply/row_matrixP=> j; pose jR := invm injhR (row j 1%:M).
+have RjR: jR \in [~: W, Ai1].
+ by rewrite -(im_invm injhR) mem_morphim /= ?im_hR ?inE.
+rewrite rowE -row1 mulmx_sumr raddf0 -/jR.
+have /subsetP nRA1: Ai1 \subset 'N([~: W, Ai1]) by rewrite commg_normr.
+transitivity (\sum_(a1 in Ai1) hR (jR ^ a1)).
+ rewrite {1}[Ai1 in rhs in _ = rhs]morphimEsub /= ?im_restrm ?imfG1 //.
+ rewrite big_imset /=; last first.
+ apply: sub_in2 (injmP (injm_invm injfG)); apply/subsetP.
+ by rewrite /= im_restrm imfG1.
+ apply: eq_bigr => a /AfG' A1a.
+ have RjRa: jR ^ fG' a \in [~: W, Ai1] by rewrite memJ_norm ?nRA1.
+ rewrite !mulmxA defPu defPl mul_rV_lin1 -/hW [hW]lock /= -lock.
+ rewrite subgK ?(subsetP sAiG1) // -DfW' !fW'K ?(subsetP sRW) //.
+ by rewrite [hW _]dprodmEl // DfUl'.
+have [nf [fj Rfj ->]] := gen_prodgP RjR.
+transitivity (\sum_(a1 in Ai1) (\prod_i1 hR (fj i1 ^ a1))%g).
+ apply: eq_bigr => a1 Aa1; rewrite conjg_prod morph_prod // => i1 _.
+ by rewrite memJ_norm ?mem_gen ?nRA1.
+rewrite exchange_big big1 //= => i1 _; have /imset2P[w a1 Ww Aa1 ->] := Rfj i1.
+apply: (addrI (\sum_(a2 in Ai1) hR [~ w, a2])).
+rewrite addr0 {2}(reindex_inj (mulgI a1)) -big_split /=.
+apply: eq_big => [a2 | a2 Aa2]; first by rewrite groupMl.
+by rewrite commgMJ [rhs in _ = rhs]morphM ?memJ_norm ?nRA1 ?mem_commg ?groupM.
+Qed.